From 94f4ade387013a2e3fe8d1042fbd152088ce1daa Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 3 Oct 2016 14:46:58 +0200 Subject: Introduction of the Ltac2 plugin. For now, it is only a simple mini-ML whose effects are the proofview monad. There is no facility to manipulate terms nor any hardwired tactic. Pattern-matching is restricted to superficial constructors, the language is lacking a lot of user-friendly interfaces, the grammar is crappy, and much more. --- Ltac2.v | 42 +++ g_ltac2.ml4 | 206 ++++++++++++ ltac2_plugin.mlpack | 6 + tac2core.ml | 119 +++++++ tac2core.mli | 34 ++ tac2entries.ml | 321 ++++++++++++++++++ tac2entries.mli | 33 ++ tac2env.ml | 139 ++++++++ tac2env.mli | 58 ++++ tac2expr.mli | 136 ++++++++ tac2intern.ml | 921 ++++++++++++++++++++++++++++++++++++++++++++++++++++ tac2intern.mli | 30 ++ tac2interp.ml | 108 ++++++ tac2interp.mli | 19 ++ vo.itarget | 1 + 15 files changed, 2173 insertions(+) create mode 100644 Ltac2.v create mode 100644 g_ltac2.ml4 create mode 100644 ltac2_plugin.mlpack create mode 100644 tac2core.ml create mode 100644 tac2core.mli create mode 100644 tac2entries.ml create mode 100644 tac2entries.mli create mode 100644 tac2env.ml create mode 100644 tac2env.mli create mode 100644 tac2expr.mli create mode 100644 tac2intern.ml create mode 100644 tac2intern.mli create mode 100644 tac2interp.ml create mode 100644 tac2interp.mli create mode 100644 vo.itarget diff --git a/Ltac2.v b/Ltac2.v new file mode 100644 index 0000000000..a952524e71 --- /dev/null +++ b/Ltac2.v @@ -0,0 +1,42 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* unit := "ltac2" "print". +Ltac2 @ external of_string : string -> message := "ltac2" "message_of_string". +Ltac2 @ external of_int : int -> message := "ltac2" "message_of_int". + +End Message. + +Module Array. + +Ltac2 @external make : int -> 'a -> ('a) array := "ltac2" "array_make". +Ltac2 @external length : ('a) array -> int := "ltac2" "array_length". +Ltac2 @external get : ('a) array -> int -> 'a := "ltac2" "array_get". +Ltac2 @external set : ('a) array -> int -> 'a -> unit := "ltac2" "array_set". + +End Array. diff --git a/g_ltac2.ml4 b/g_ltac2.ml4 new file mode 100644 index 0000000000..349220f9de --- /dev/null +++ b/g_ltac2.ml4 @@ -0,0 +1,206 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* CPatRef (!@loc, id, pl) + | id = Prim.qualid -> CPatRef (!@loc, id, []) ] + | "0" + [ "_" -> CPatAny (!@loc) + | "()" -> CPatTup (!@loc, []) + | id = Prim.qualid -> CPatRef (!@loc, id, []) + | "("; pl = LIST0 tac2pat LEVEL "1" SEP ","; ")" -> CPatTup (!@loc, pl) ] + ] + ; + tac2expr: + [ "5" + [ "fun"; it = LIST1 input_fun ; "=>"; body = tac2expr LEVEL "5" -> CTacFun (!@loc, it, body) + | "let"; isrec = rec_flag; + lc = LIST1 let_clause SEP "with"; "in"; + e = tac2expr LEVEL "5" -> CTacLet (!@loc, isrec, lc, e) + | "match"; e = tac2expr LEVEL "5"; "with"; bl = branches ;"end" -> + CTacCse (!@loc, e, bl) + ] + | "2" LEFTA + [ e1 = tac2expr; ";"; e2 = tac2expr -> CTacSeq (!@loc, e1, e2) ] + | "1" LEFTA + [ e = tac2expr; el = LIST1 tac2expr LEVEL "0" -> CTacApp (!@loc, e, el) + | e0 = tac2expr; ","; el = LIST1 tac2expr LEVEL "0" SEP "," -> CTacTup (!@loc, e0 :: el) ] + | "0" + [ "("; a = tac2expr LEVEL "5"; ")" -> a + | "("; a = tac2expr; ":"; t = tac2type; ")" -> CTacCnv (!@loc, a, t) + | "()" -> CTacTup (!@loc, []) + | "("; ")" -> CTacTup (!@loc, []) + | "["; a = LIST0 tac2expr LEVEL "1" SEP ";"; "]" -> CTacLst (!@loc, a) + | a = tactic_atom -> a ] + ] + ; + branches: + [ [ -> [] + | "|"; bl = LIST1 branch SEP "|" -> bl + | bl = LIST1 branch SEP "|" -> bl ] + ] + ; + branch: + [ [ pat = tac2pat LEVEL "1"; "=>"; e = tac2expr LEVEL "5" -> (pat, e) ] ] + ; + rec_flag: + [ [ IDENT "rec" -> true + | -> false ] ] + ; + typ_param: + [ [ "'"; id = Prim.ident -> id ] ] + ; + tactic_atom: + [ [ n = Prim.integer -> CTacAtm (!@loc, AtmInt n) + | s = Prim.string -> CTacAtm (!@loc, AtmStr s) + | id = Prim.qualid -> CTacRef id + ] ] + ; + let_clause: + [ [ id = binder; ":="; te = tac2expr -> + (id, None, te) + | id = binder; args = LIST1 input_fun; ":="; te = tac2expr -> + (id, None, CTacFun (!@loc, args, te)) ] ] + ; + tac2type: + [ "5" RIGHTA + [ t1 = tac2type; "->"; t2 = tac2type -> CTypArrow (!@loc, t1, t2) ] + | "2" + [ t = tac2type; "*"; tl = LIST1 tac2type SEP "*" -> CTypTuple (!@loc, t :: tl) ] + | "1" + [ "("; p = LIST1 tac2type LEVEL "5" SEP ","; ")"; qid = Prim.qualid -> CTypRef (!@loc, qid, p) ] + | "0" + [ "("; t = tac2type; ")" -> t + | id = typ_param -> CTypVar (!@loc, Name id) + | "_" -> CTypVar (!@loc, Anonymous) + | qid = Prim.qualid -> CTypRef (!@loc, qid, []) ] + ]; + locident: + [ [ id = Prim.ident -> (!@loc, id) ] ] + ; + binder: + [ [ "_" -> (!@loc, Anonymous) + | l = Prim.ident -> (!@loc, Name l) ] ] + ; + input_fun: + [ [ b = binder -> (b, None) + | "("; b = binder; ":"; t = tac2type; ")" -> (b, Some t) ] ] + ; + tac2def_body: + [ [ name = binder; it = LIST0 input_fun; ":="; e = tac2expr -> + let e = if List.is_empty it then e else CTacFun (!@loc, it, e) in + (name, e) + ] ] + ; + tac2def_val: + [ [ isrec = rec_flag; l = LIST1 tac2def_body SEP "with" -> + StrVal (isrec, l) + ] ] + ; + tac2typ_knd: + [ [ t = tac2type -> CTydDef (Some t) + | t = tac2alg_type -> CTydAlg t ] ] + ; + tac2alg_type: + [ [ -> [] + | "|"; bl = LIST1 tac2alg_constructor SEP "|" -> bl ] ] + ; + tac2alg_constructor: + [ [ c = Prim.ident -> (c, []) + | c = Prim.ident; "("; args = LIST0 tac2type SEP ","; ")"-> (c, args) ] ] + ; + tac2typ_prm: + [ [ -> [] + | id = typ_param -> [!@loc, id] + | "("; ids = LIST1 [ id = typ_param -> (!@loc, id) ] SEP "," ;")" -> ids + ] ] + ; + tac2typ_def: + [ [ prm = tac2typ_prm; id = locident; ":="; e = tac2typ_knd -> + (id, (prm, e)) + | prm = tac2typ_prm; id = locident -> (id, (prm, CTydDef None)) + ] ] + ; + tac2def_typ: + [ [ "Type"; isrec = rec_flag; l = LIST1 tac2typ_def SEP "with" -> + StrTyp (isrec, l) + ] ] + ; + tac2def_ext: + [ [ "@"; IDENT "external"; id = locident; ":"; t = tac2type LEVEL "5"; ":="; + plugin = Prim.string; name = Prim.string -> + let ml = { mltac_plugin = plugin; mltac_tactic = name } in + StrPrm (id, t, ml) + ] ] + ; +END + +let pr_ltac2entry _ = mt () (** FIXME *) +let pr_ltac2expr _ = mt () (** FIXME *) + +VERNAC ARGUMENT EXTEND ltac2_entry +PRINTED BY pr_ltac2entry +| [ tac2def_val(v) ] -> [ v ] +| [ tac2def_typ(t) ] -> [ t ] +| [ tac2def_ext(e) ] -> [ e ] +END + +VERNAC COMMAND EXTEND VernacDeclareTactic2Definition CLASSIFIED AS SIDEFF +| [ "Ltac2" ltac2_entry(e) ] -> [ + let local = Locality.LocalityFixme.consume () in + Tac2entries.register_struct ?local e + ] +END + +let _ = + let mode = { + Proof_global.name = "Ltac2"; + set = (fun () -> set_command_entry tac2mode); + reset = (fun () -> set_command_entry Pcoq.Vernac_.noedit_mode); + } in + Proof_global.register_proof_mode mode + +VERNAC ARGUMENT EXTEND ltac2_expr +PRINTED BY pr_ltac2expr +| [ tac2expr(e) ] -> [ e ] +END + +open G_ltac +open Vernac_classifier + +VERNAC tac2mode EXTEND VernacLtac2 +| [ - ltac2_expr(t) ltac_use_default(default) ] => + [ classify_as_proofstep ] -> [ +(* let g = Option.default (Proof_global.get_default_goal_selector ()) g in *) + Tac2entries.call ~default t + ] +END + +open Stdarg + +VERNAC COMMAND EXTEND Ltac2Print CLASSIFIED AS SIDEFF +| [ "Print" "Ltac2" reference(tac) ] -> [ Tac2entries.print_ltac tac ] +END diff --git a/ltac2_plugin.mlpack b/ltac2_plugin.mlpack new file mode 100644 index 0000000000..561bd0eb0a --- /dev/null +++ b/ltac2_plugin.mlpack @@ -0,0 +1,6 @@ +Tac2env +Tac2intern +Tac2interp +Tac2entries +Tac2core +G_ltac2 diff --git a/tac2core.ml b/tac2core.ml new file mode 100644 index 0000000000..af4124e647 --- /dev/null +++ b/tac2core.ml @@ -0,0 +1,119 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* () +| _ -> assert false + +let rec of_list = function +| [] -> v_nil +| x :: l -> v_cons x (of_list l) + +let rec to_list = function +| ValInt 0 -> [] +| ValBlk (0, [|v; vl|]) -> v :: to_list vl +| _ -> assert false + +end + +let extract_val (type a) (tag : a Val.typ) (Val.Dyn (tag', v)) : a = +match Val.eq tag tag' with +| None -> assert false +| Some Refl -> v + +let val_pp = Val.create "ltac2:pp" +let get_pp v = extract_val val_pp v + +(** Helper functions *) + +let return x = Proofview.tclUNIT x +let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } + +let wrap f = + return () >>= fun () -> return (f ()) + +let wrap_unit f = + return () >>= fun () -> f (); return v_unit + +(** Primitives *) + +let prm_print : ml_tactic = function +| [ValExt pp] -> wrap_unit (fun () -> Feedback.msg_notice (get_pp pp)) +| _ -> assert false + +let prm_message_of_int : ml_tactic = function +| [ValInt s] -> return (ValExt (Val.Dyn (val_pp, int s))) +| _ -> assert false + +let prm_message_of_string : ml_tactic = function +| [ValStr s] -> return (ValExt (Val.Dyn (val_pp, str s))) +| _ -> assert false + +let prm_array_make : ml_tactic = function +| [ValInt n; x] -> wrap (fun () -> ValBlk (0, Array.make n x)) +| _ -> assert false + +let prm_array_length : ml_tactic = function +| [ValBlk (_, v)] -> return (ValInt (Array.length v)) +| _ -> assert false + +let prm_array_set : ml_tactic = function +| [ValBlk (_, v); ValInt n; x] -> wrap_unit (fun () -> v.(n) <- x) +| _ -> assert false + +let prm_array_get : ml_tactic = function +| [ValBlk (_, v); ValInt n] -> wrap (fun () -> v.(n)) +| _ -> assert false + +(** Registering *) + +let () = Tac2env.define_primitive (pname "print") prm_print +let () = Tac2env.define_primitive (pname "message_of_string") prm_message_of_string +let () = Tac2env.define_primitive (pname "message_of_int") prm_message_of_int + +let () = Tac2env.define_primitive (pname "array_make") prm_array_make +let () = Tac2env.define_primitive (pname "array_length") prm_array_length +let () = Tac2env.define_primitive (pname "array_get") prm_array_get +let () = Tac2env.define_primitive (pname "array_set") prm_array_set diff --git a/tac2core.mli b/tac2core.mli new file mode 100644 index 0000000000..14bde483c1 --- /dev/null +++ b/tac2core.mli @@ -0,0 +1,34 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* valexpr +val to_unit : valexpr -> unit + +val of_list : valexpr list -> valexpr +val to_list : valexpr -> valexpr list + +end diff --git a/tac2entries.ml b/tac2entries.ml new file mode 100644 index 0000000000..7572270ab3 --- /dev/null +++ b/tac2entries.ml @@ -0,0 +1,321 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* obj = + declare_object {(default_object "TAC2-DEFINITION") with + cache_function = cache_tacdef; + load_function = load_tacdef; + open_function = open_tacdef; + subst_function = subst_tacdef; + classify_function = classify_tacdef} + +type typdef = { + typdef_local : bool; + typdef_expr : glb_quant_typedef; +} + +let change_kn_label kn id = + let (mp, dp, _) = KerName.repr kn in + KerName.make mp dp (Label.of_id id) + +let change_sp_label sp id = + let (dp, _) = Libnames.repr_path sp in + Libnames.make_path dp id + +let push_typedef visibility sp kn (_, def) = match def with +| GTydDef _ -> + Tac2env.push_type visibility sp kn +| GTydAlg cstrs -> + (** Register constructors *) + let iter (c, _) = + let spc = change_sp_label sp c in + let knc = change_kn_label kn c in + Tac2env.push_ltac visibility spc knc + in + Tac2env.push_type visibility sp kn; + List.iter iter cstrs +| GTydRec _ -> + assert false (** FIXME *) + +let next i = + let ans = !i in + let () = incr i in + ans + +let dummy_var i = Id.of_string (Printf.sprintf "_%i" i) + +let define_typedef kn (params, def as qdef) = match def with +| GTydDef _ -> + Tac2env.define_type kn qdef +| GTydAlg cstrs -> + (** Define constructors *) + let constant = ref 0 in + let nonconstant = ref 0 in + let iter (c, args) = + let knc = change_kn_label kn c in + let tag = if List.is_empty args then next constant else next nonconstant in + let ids = List.mapi (fun i _ -> dummy_var i) args in + let c = GTacCst (kn, tag, List.map (fun id -> GTacVar id) ids) in + let c = + if List.is_empty args then c + else GTacFun (List.map (fun id -> Name id) ids, c) + in + let fold arg tpe = GTypArrow (arg, tpe) in + let cT = GTypRef (kn, List.init params (fun i -> GTypVar i)) in + let cT = List.fold_right fold args cT in + let data = { + Tac2env.cdata_type = kn; + cdata_args = params, args; + cdata_indx = tag; + } in + Tac2env.define_constructor knc data; + Tac2env.define_global knc (c, (params, cT)) + in + Tac2env.define_type kn qdef; + List.iter iter cstrs +| GTydRec _ -> + assert false (** FIXME *) + +let perform_typdef vs ((sp, kn), def) = + let () = if not def.typdef_local then push_typedef vs sp kn def.typdef_expr in + define_typedef kn def.typdef_expr + +let load_typdef i obj = perform_typdef (Until i) obj +let open_typdef i obj = perform_typdef (Exactly i) obj + +let cache_typdef ((sp, kn), def) = + let () = push_typedef (Until 1) sp kn def.typdef_expr in + define_typedef kn def.typdef_expr + +let subst_typdef (subst, def) = + let expr' = subst_quant_typedef subst def.typdef_expr in + if expr' == def.typdef_expr then def else { def with typdef_expr = expr' } + +let classify_typdef o = Substitute o + +let inTypDef : typdef -> obj = + declare_object {(default_object "TAC2-TYPE-DEFINITION") with + cache_function = cache_typdef; + load_function = load_typdef; + open_function = open_typdef; + subst_function = subst_typdef; + classify_function = classify_typdef} + +let register_ltac ?(local = false) isrec tactics = + if isrec then + let map (na, e) = (na, None, e) in + let bindings = List.map map tactics in + let map ((loc, na), e) = match na with + | Anonymous -> None + | Name id -> + let qid = Libnames.qualid_of_ident id in + let e = CTacLet (Loc.ghost, true, bindings, CTacRef (loc, qid)) in + let (e, t) = intern e in + let e = match e with + | GTacLet (true, _, e) -> assert false + | _ -> assert false + in + Some (e, t) + in + let tactics = List.map map tactics in + assert false (** FIXME *) + else + let map ((loc, na), e) = + let (e, t) = intern e in + let () = + if not (is_value e) then + user_err ~loc (str "Tactic definition must be a syntactical value") + in + let id = match na with + | Anonymous -> + user_err ~loc (str "Tactic definition must have a name") + | Name id -> id + in + let kn = Lib.make_kn id in + let exists = + try let _ = Tac2env.interp_global kn in true with Not_found -> false + in + let () = + if exists then + user_err ~loc (str "Tactic " ++ Nameops.pr_id id ++ str " already exists") + in + (id, e, t) + in + let defs = List.map map tactics in + let iter (id, e, t) = + let def = { + tacdef_local = local; + tacdef_expr = e; + tacdef_type = t; + } in + ignore (Lib.add_leaf id (inTacDef def)) + in + List.iter iter defs + +let register_type ?(local = false) isrec types = + let same_name ((_, id1), _) ((_, id2), _) = Id.equal id1 id2 in + let () = match List.duplicates same_name types with + | [] -> () + | ((loc, id), _) :: _ -> + user_err ~loc (str "Multiple definition of the type name " ++ Id.print id) + in + let check ((loc, id), (params, def)) = + let same_name (_, id1) (_, id2) = Id.equal id1 id2 in + let () = match List.duplicates same_name params with + | [] -> () + | (loc, id) :: _ -> + user_err ~loc (str "The type parameter " ++ Id.print id ++ + str " occurs several times") + in + match def with + | CTydDef _ -> + if isrec then + user_err ~loc (str "The type abbreviation " ++ Id.print id ++ + str " cannot be recursive") + | CTydAlg _ -> () (** FIXME *) + | CTydRec _ -> assert false (** FIXME *) + in + let () = List.iter check types in + let self = + if isrec then + let fold accu ((_, id), (params, _)) = + Id.Map.add id (Lib.make_kn id, List.length params) accu + in + List.fold_left fold Id.Map.empty types + else Id.Map.empty + in + let map ((_, id), def) = + let typdef = { + typdef_local = local; + typdef_expr = intern_typedef self def; + } in + (id, typdef) + in + let types = List.map map types in + let iter (id, def) = ignore (Lib.add_leaf id (inTypDef def)) in + List.iter iter types + +let register_primitive ?(local = false) (loc, id) t ml = + let t = intern_open_type t in + let rec count_arrow = function + | GTypArrow (_, t) -> 1 + count_arrow t + | _ -> 0 + in + let arrows = count_arrow (snd t) in + let () = if Int.equal arrows 0 then + user_err ~loc (str "External tactic must have at least one argument") in + let () = + try let _ = Tac2env.interp_primitive ml in () with Not_found -> + user_err ~loc (str "Unregistered primitive " ++ + quote (str ml.mltac_plugin) ++ spc () ++ quote (str ml.mltac_tactic)) + in + let init i = Id.of_string (Printf.sprintf "x%i" i) in + let names = List.init arrows init in + let bnd = List.map (fun id -> Name id) names in + let arg = List.map (fun id -> GTacVar id) names in + let e = GTacFun (bnd, GTacPrm (ml, arg)) in + let def = { + tacdef_local = local; + tacdef_expr = e; + tacdef_type = t; + } in + ignore (Lib.add_leaf id (inTacDef def)) + +let register_struct ?local str = match str with +| StrVal (isrec, e) -> register_ltac ?local isrec e +| StrTyp (isrec, t) -> register_type ?local isrec t +| StrPrm (id, t, ml) -> register_primitive ?local id t ml + +(** Printing *) + +let print_ltac ref = + let (loc, qid) = qualid_of_reference ref in + let kn = + try Tac2env.locate_ltac qid + with Not_found -> user_err ~loc (str "Unknown tactic " ++ pr_qualid qid) + in + let (_, (_, t)) = Tac2env.interp_global kn in + Feedback.msg_notice (pr_qualid qid ++ spc () ++ str ":" ++ spc () ++ pr_glbtype t) + +(** Calling tactics *) + +let solve default tac = + let status = Proof_global.with_current_proof begin fun etac p -> + let with_end_tac = if default then Some etac else None in + let (p, status) = Pfedit.solve SelectAll None tac ?with_end_tac p in + (* in case a strict subtree was completed, + go back to the top of the prooftree *) + let p = Proof.maximal_unfocus Vernacentries.command_focus p in + p, status + end in + if not status then Feedback.feedback Feedback.AddedAxiom + +let call ~default e = + let loc = loc_of_tacexpr e in + let (e, (_, t)) = intern e in + let () = check_unit ~loc t in + let tac = Tac2interp.interp Id.Map.empty e in + solve default (Proofview.tclIGNORE tac) + +(** Primitive algebraic types than can't be defined Coq-side *) + +let register_prim_alg name params def = + let id = Id.of_string name in + let def = List.map (fun (cstr, tpe) -> (Id.of_string_soft cstr, tpe)) def in + let def = (params, GTydAlg def) in + let def = { typdef_local = false; typdef_expr = def } in + ignore (Lib.add_leaf id (inTypDef def)) + +let coq_prefix = DirPath.make (List.map Id.of_string ["Ltac2"; "ltac2"; "Coq"]) +let coq_def n = KerName.make2 (MPfile coq_prefix) (Label.make n) + +let t_list = coq_def "list" + +let _ = Mltop.declare_cache_obj begin fun () -> + register_prim_alg "unit" 0 ["()", []]; + register_prim_alg "list" 1 [ + ("[]", []); + ("::", [GTypVar 0; GTypRef (t_list, [GTypVar 0])]); + ]; +end "ltac2_plugin" diff --git a/tac2entries.mli b/tac2entries.mli new file mode 100644 index 0000000000..9c5d0a15fd --- /dev/null +++ b/tac2entries.mli @@ -0,0 +1,33 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* rec_flag -> + (Name.t located * raw_tacexpr) list -> unit + +val register_type : ?local:bool -> rec_flag -> + (Id.t located * raw_quant_typedef) list -> unit + +val register_primitive : ?local:bool -> + Id.t located -> raw_typexpr -> ml_tactic_name -> unit + +val register_struct : ?local:bool -> strexpr -> unit + +(** {5 Inspecting} *) + +val print_ltac : Libnames.reference -> unit + +(** {5 Eval loop} *) + +(** Evaluate a tactic expression in the current environment *) +val call : default:bool -> raw_tacexpr -> unit diff --git a/tac2env.ml b/tac2env.ml new file mode 100644 index 0000000000..d17e657516 --- /dev/null +++ b/tac2env.ml @@ -0,0 +1,139 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* ValInt n +| GTacRef kn -> + let (e, _) = + try KNmap.find kn ltac_state.contents.ltac_tactics + with Not_found -> assert false + in + eval_pure e +| GTacFun (na, e) -> + ValCls { clos_env = Id.Map.empty; clos_var = na; clos_exp = e } +| GTacTup el -> ValBlk (0, Array.map_of_list eval_pure el) +| GTacCst (_, n, []) -> ValInt n +| GTacCst (_, n, el) -> ValBlk (n, Array.map_of_list eval_pure el) +| GTacAtm (AtmStr _) | GTacArr _ | GTacLet _ | GTacVar _ +| GTacApp _ | GTacCse _ | GTacPrm _ -> + anomaly (Pp.str "Term is not a syntactical value") + +let define_global kn e = + let state = !ltac_state in + ltac_state := { state with ltac_tactics = KNmap.add kn e state.ltac_tactics } + +let interp_global kn = + let (e, t) = KNmap.find kn ltac_state.contents.ltac_tactics in + (eval_pure e, t) + +let define_constructor kn t = + let state = !ltac_state in + ltac_state := { state with ltac_constructors = KNmap.add kn t state.ltac_constructors } + +let interp_constructor kn = KNmap.find kn ltac_state.contents.ltac_constructors + +let define_type kn e = + let state = !ltac_state in + ltac_state := { state with ltac_types = KNmap.add kn e state.ltac_types } + +let interp_type kn = KNmap.find kn ltac_state.contents.ltac_types + +module ML = +struct + type t = ml_tactic_name + let compare n1 n2 = + let c = String.compare n1.mltac_plugin n2.mltac_plugin in + if Int.equal c 0 then String.compare n1.mltac_tactic n2.mltac_tactic + else c +end + +module MLMap = Map.Make(ML) + +let primitive_map = ref MLMap.empty + +let define_primitive name f = primitive_map := MLMap.add name f !primitive_map +let interp_primitive name = MLMap.find name !primitive_map + +(** Name management *) + +module FullPath = +struct + type t = full_path + let equal = eq_full_path + let to_string = string_of_path + let repr sp = + let dir,id = repr_path sp in + id, (DirPath.repr dir) +end + +module KnTab = Nametab.Make(FullPath)(KerName) + +type nametab = { + tab_ltac : KnTab.t; + tab_type : KnTab.t; +} + +let empty_nametab = { tab_ltac = KnTab.empty; tab_type = KnTab.empty } + +let nametab = Summary.ref empty_nametab ~name:"ltac2-nametab" + +let push_ltac vis sp kn = + let tab = !nametab in + nametab := { tab with tab_ltac = KnTab.push vis sp kn tab.tab_ltac } + +let locate_ltac qid = + let tab = !nametab in + KnTab.locate qid tab.tab_ltac + +let locate_extended_all_ltac qid = + let tab = !nametab in + KnTab.find_prefixes qid tab.tab_ltac + +let push_type vis sp kn = + let tab = !nametab in + nametab := { tab with tab_type = KnTab.push vis sp kn tab.tab_type } + +let locate_type qid = + let tab = !nametab in + KnTab.locate qid tab.tab_type + +let locate_extended_all_type qid = + let tab = !nametab in + KnTab.find_prefixes qid tab.tab_type diff --git a/tac2env.mli b/tac2env.mli new file mode 100644 index 0000000000..efabdb5466 --- /dev/null +++ b/tac2env.mli @@ -0,0 +1,58 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* (glb_tacexpr * type_scheme) -> unit +val interp_global : ltac_constant -> (valexpr * type_scheme) + +(** {5 Toplevel definition of types} *) + +val define_type : type_constant -> glb_quant_typedef -> unit +val interp_type : type_constant -> glb_quant_typedef + +(** {5 Toplevel definition of algebraic constructors} *) + +type constructor_data = { + cdata_type : type_constant; + cdata_args : int * int glb_typexpr list; + cdata_indx : int; +} + +val define_constructor : ltac_constructor -> constructor_data -> unit +val interp_constructor : ltac_constructor -> constructor_data + +(** {5 Name management} *) + +val push_ltac : visibility -> full_path -> ltac_constant -> unit +val locate_ltac : qualid -> ltac_constant +val locate_extended_all_ltac : qualid -> ltac_constant list + +val push_type : visibility -> full_path -> type_constant -> unit +val locate_type : qualid -> type_constant +val locate_extended_all_type : qualid -> type_constant list + +(** {5 Toplevel definitions of ML tactics} *) + +(** This state is not part of the summary, contrarily to the ones above. It is + intended to be used from ML plugins to register ML-side functions. *) + +val define_primitive : ml_tactic_name -> ml_tactic -> unit +val interp_primitive : ml_tactic_name -> ml_tactic diff --git a/tac2expr.mli b/tac2expr.mli new file mode 100644 index 0000000000..445c69aa23 --- /dev/null +++ b/tac2expr.mli @@ -0,0 +1,136 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* valexpr Proofview.tactic diff --git a/tac2intern.ml b/tac2intern.ml new file mode 100644 index 0000000000..b1b35b0787 --- /dev/null +++ b/tac2intern.ml @@ -0,0 +1,921 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* elt -> bool +val create : unit -> 'a t +val fresh : 'a t -> elt +val size : 'a t -> int +val find : elt -> 'a t -> (elt * 'a option) +val union : elt -> elt -> 'a t -> unit +val set : elt -> 'a -> 'a t -> unit +module Map : +sig + type key = elt + type +'a t + val empty : 'a t + val add : key -> 'a -> 'a t -> 'a t + val mem : key -> 'a t -> bool + val find : key -> 'a t -> 'a +end +end += +struct +type elt = int +let equal = Int.equal +module Map = Int.Map + +type 'a node = +| Canon of int * 'a option +| Equiv of elt + +type 'a t = { + mutable uf_data : 'a node array; + mutable uf_size : int; +} + +let size p = p.uf_size + +let resize p = + if Int.equal (Array.length p.uf_data) p.uf_size then begin + let nsize = 2 * p.uf_size + 1 in + let v = Array.make nsize (Equiv 0) in + Array.blit p.uf_data 0 v 0 (Array.length p.uf_data); + p.uf_data <- v; + end + +let create () = { uf_data = [||]; uf_size = 0 } + +let fresh p = + resize p; + let n = p.uf_size in + p.uf_data.(n) <- (Canon (1, None)); + p.uf_size <- n + 1; + n + +let rec lookup n p = + let node = Array.get p.uf_data n in + match node with + | Canon (size, v) -> n, size, v + | Equiv y -> + let ((z, _, _) as res) = lookup y p in + if not (Int.equal z y) then Array.set p.uf_data n (Equiv z); + res + +let find n p = + let (x, _, v) = lookup n p in (x, v) + +let union x y p = + let ((x, size1, _) as xcan) = lookup x p in + let ((y, size2, _) as ycan) = lookup y p in + let xcan, ycan = if size1 < size2 then xcan, ycan else ycan, xcan in + let x, _, xnode = xcan in + let y, _, ynode = ycan in + assert (Option.is_empty xnode); + assert (Option.is_empty ynode); + p.uf_data.(x) <- Equiv y; + p.uf_data.(y) <- Canon (size1 + size2, None) + +let set x v p = + let (x, s, v') = lookup x p in + assert (Option.is_empty v'); + p.uf_data.(x) <- Canon (s, Some v) + +end + +type mix_var = +| GVar of UF.elt +| LVar of int + +type mix_type_scheme = int * mix_var glb_typexpr + +type environment = { + env_var : mix_type_scheme Id.Map.t; + (** Type schemes of bound variables *) + env_cst : UF.elt glb_typexpr UF.t; + (** Unification state *) + env_als : UF.elt Id.Map.t ref; + (** Map user-facing type variables to unification variables *) + env_opn : bool; + (** Accept unbound type variables *) + env_rec : (KerName.t * int) Id.Map.t; + (** Recursive type definitions *) +} + +let empty_env () = { + env_var = Id.Map.empty; + env_cst = UF.create (); + env_als = ref Id.Map.empty; + env_opn = true; + env_rec = Id.Map.empty; +} + +let fresh_id env = UF.fresh env.env_cst + +let get_alias (loc, id) env = + try Id.Map.find id env.env_als.contents + with Not_found -> + if env.env_opn then + let n = fresh_id env in + let () = env.env_als := Id.Map.add id n env.env_als.contents in + n + else user_err ~loc (str "Unbound type parameter " ++ Id.print id) + +let push_name id t env = match id with +| Anonymous -> env +| Name id -> { env with env_var = Id.Map.add id t env.env_var } + +let loc_of_tacexpr = function +| CTacAtm (loc, _) -> loc +| CTacRef (loc, _) -> loc +| CTacFun (loc, _, _) -> loc +| CTacApp (loc, _, _) -> loc +| CTacLet (loc, _, _, _) -> loc +| CTacTup (loc, _) -> loc +| CTacArr (loc, _) -> loc +| CTacLst (loc, _) -> loc +| CTacCnv (loc, _, _) -> loc +| CTacSeq (loc, _, _) -> loc +| CTacCse (loc, _, _) -> loc + +let loc_of_patexpr = function +| CPatAny loc -> loc +| CPatRef (loc, _, _) -> loc +| CPatTup (loc, _) -> loc + +let rec subst_type subst (t : 'a glb_typexpr) = match t with +| GTypVar id -> subst id +| GTypArrow (t1, t2) -> GTypArrow (subst_type subst t1, subst_type subst t2) +| GTypTuple tl -> GTypTuple (List.map (fun t -> subst_type subst t) tl) +| GTypRef (qid, args) -> + GTypRef (qid, List.map (fun t -> subst_type subst t) args) + +let rec intern_type env (t : raw_typexpr) : UF.elt glb_typexpr = match t with +| CTypVar (loc, Name id) -> GTypVar (get_alias (loc, id) env) +| CTypVar (_, Anonymous) -> GTypVar (fresh_id env) +| CTypRef (_, (loc, qid), args) -> + let (dp, id) = repr_qualid qid in + let (kn, nparams) = + if DirPath.is_empty dp && Id.Map.mem id env.env_rec then + Id.Map.find id env.env_rec + else + let kn = + try Tac2env.locate_type qid + with Not_found -> + user_err ~loc (str "Unbound type constructor " ++ pr_qualid qid) + in + let (nparams, _) = Tac2env.interp_type kn in + (kn, nparams) + in + let nargs = List.length args in + let () = + if not (Int.equal nparams nargs) then + user_err ~loc (strbrk "The type constructor " ++ pr_qualid qid ++ + strbrk " expects " ++ int nparams ++ strbrk " argument(s), but is here \ + applied to " ++ int nargs ++ strbrk "argument(s)") + in + GTypRef (kn, List.map (fun t -> intern_type env t) args) +| CTypArrow (loc, t1, t2) -> GTypArrow (intern_type env t1, intern_type env t2) +| CTypTuple (loc, tl) -> GTypTuple (List.map (fun t -> intern_type env t) tl) + +let fresh_type_scheme env (t : type_scheme) : UF.elt glb_typexpr = + let (n, t) = t in + let subst = Array.init n (fun _ -> fresh_id env) in + let substf i = GTypVar subst.(i) in + subst_type substf t + +let fresh_mix_type_scheme env (t : mix_type_scheme) : UF.elt glb_typexpr = + let (n, t) = t in + let subst = Array.init n (fun _ -> fresh_id env) in + let substf = function + | LVar i -> GTypVar subst.(i) + | GVar n -> GTypVar n + in + subst_type substf t + +let fresh_reftype env (kn : KerName.t) = + let (n, _) = Tac2env.interp_type kn in + let subst = Array.init n (fun _ -> fresh_id env) in + let t = GTypRef (kn, Array.map_to_list (fun i -> GTypVar i) subst) in + (subst, t) + +(** First-order unification algorithm *) + +let is_unfoldable kn = match snd (Tac2env.interp_type kn) with +| GTydDef (Some _) -> true +| GTydDef None | GTydAlg _ | GTydRec _ -> false + +let unfold env kn args = + let (nparams, def) = Tac2env.interp_type kn in + let def = match def with + | GTydDef (Some t) -> t + | _ -> assert false + in + let args = Array.of_list args in + let subst n = args.(n) in + subst_type subst def + +(** View function, allows to ensure head normal forms *) +let rec kind env t = match t with +| GTypVar id -> + let (id, v) = UF.find id env.env_cst in + begin match v with + | None -> GTypVar id + | Some t -> kind env t + end +| GTypRef (kn, tl) -> + if is_unfoldable kn then kind env (unfold env kn tl) else t +| GTypArrow _ | GTypTuple _ -> t + +exception Occur + +let rec occur_check env id t = match kind env t with +| GTypVar id' -> if UF.equal id id' then raise Occur +| GTypArrow (t1, t2) -> + let () = occur_check env id t1 in + occur_check env id t2 +| GTypTuple tl -> + List.iter (fun t -> occur_check env id t) tl +| GTypRef (kn, tl) -> + List.iter (fun t -> occur_check env id t) tl + +exception CannotUnify of UF.elt glb_typexpr * UF.elt glb_typexpr + +let unify_var env id t = match kind env t with +| GTypVar id' -> + if not (UF.equal id id') then UF.union id id' env.env_cst +| GTypArrow _ | GTypRef _ | GTypTuple _ -> + try + let () = occur_check env id t in + UF.set id t env.env_cst + with Occur -> raise (CannotUnify (GTypVar id, t)) + +let rec unify env t1 t2 = match kind env t1, kind env t2 with +| GTypVar id, t | t, GTypVar id -> + unify_var env id t +| GTypArrow (t1, u1), GTypArrow (t2, u2) -> + let () = unify env t1 t2 in + unify env u1 u2 +| GTypTuple tl1, GTypTuple tl2 -> + if Int.equal (List.length tl1) (List.length tl2) then + List.iter2 (fun t1 t2 -> unify env t1 t2) tl1 tl2 + else raise (CannotUnify (t1, t2)) +| GTypRef (kn1, tl1), GTypRef (kn2, tl2) -> + if KerName.equal kn1 kn2 then + List.iter2 (fun t1 t2 -> unify env t1 t2) tl1 tl2 + else raise (CannotUnify (t1, t2)) +| _ -> raise (CannotUnify (t1, t2)) + +(** FIXME *) +let rec pr_glbtype = function +| GTypVar n -> str "?" +| GTypRef (kn, tl) -> + KerName.print kn ++ str "(" ++ prlist_with_sep (fun () -> str ", ") pr_glbtype tl ++ str ")" +| GTypArrow (t1, t2) -> str "Arr(" ++ pr_glbtype t1 ++ str ", " ++ pr_glbtype t2 ++ str ")" +| GTypTuple tl -> str "Tup(" ++ prlist_with_sep (fun () -> str ", ") pr_glbtype tl ++ str ")" + +let unify loc env t1 t2 = + try unify env t1 t2 + with CannotUnify (u1, u2) -> + user_err ~loc (str "This expression has type " ++ pr_glbtype t1 ++ + str " but an expression what expected of type " ++ pr_glbtype t2) + +(** Term typing *) + +let rec is_value = function +| GTacAtm (AtmInt _) | GTacVar _ | GTacRef _ | GTacFun _ -> true +| GTacAtm (AtmStr _) | GTacApp _ | GTacLet _ -> false +| GTacTup el -> List.for_all is_value el +| GTacCst (_, _, []) -> true +| GTacCst (kn, n, el) -> + (** To be a value, a constructor must be immutable *) + assert false (** TODO *) +| GTacArr _ | GTacCse _ | GTacPrm _ -> false + +let is_rec_rhs = function +| GTacFun _ -> true +| GTacAtm _ | GTacVar _ | GTacRef _ | GTacApp _ | GTacLet _ -> false +| GTacTup _ | GTacArr _ | GTacPrm _ | GTacCst _ | GTacCse _ -> false + +let rec fv_type f t accu = match t with +| GTypVar id -> f id accu +| GTypArrow (t1, t2) -> fv_type f t1 (fv_type f t2 accu) +| GTypTuple tl -> List.fold_left (fun accu t -> fv_type f t accu) accu tl +| GTypRef (kn, tl) -> List.fold_left (fun accu t -> fv_type f t accu) accu tl + +let fv_env env = + let rec f id accu = match UF.find id env.env_cst with + | id, None -> UF.Map.add id () accu + | _, Some t -> fv_type f t accu + in + let fold_var id (_, t) accu = + let fmix id accu = match id with + | LVar _ -> accu + | GVar id -> f id accu + in + fv_type fmix t accu + in + let fv_var = Id.Map.fold fold_var env.env_var UF.Map.empty in + let fold_als _ id accu = f id accu in + Id.Map.fold fold_als !(env.env_als) fv_var + +let abstract_var env (t : UF.elt glb_typexpr) : mix_type_scheme = + let fv = fv_env env in + let count = ref 0 in + let vars = ref UF.Map.empty in + let rec subst id = + let (id, t) = UF.find id env.env_cst in + match t with + | None -> + if UF.Map.mem id fv then GTypVar (GVar id) + else + begin try UF.Map.find id !vars + with Not_found -> + let n = !count in + let var = GTypVar (LVar n) in + let () = incr count in + let () = vars := UF.Map.add id var !vars in + var + end + | Some t -> subst_type subst t + in + let t = subst_type subst t in + (!count, t) + +let monomorphic (t : UF.elt glb_typexpr) : mix_type_scheme = + let subst id = GTypVar (GVar id) in + (0, subst_type subst t) + +let warn_not_unit = + CWarnings.create ~name:"not-unit" ~category:"ltac" + (fun () -> strbrk "The following expression should have type unit.") + +let warn_redundant_clause = + CWarnings.create ~name:"redundant-clause" ~category:"ltac" + (fun () -> strbrk "The following clause is redundant.") + +let check_elt_unit loc env t = + let maybe_unit = match kind env t with + | GTypVar _ -> true + | GTypArrow _ | GTypTuple _ -> false + | GTypRef (kn, _) -> KerName.equal kn t_unit + in + if not maybe_unit then warn_not_unit ~loc () + +let check_elt_empty loc env t = match kind env t with +| GTypVar _ -> + user_err ~loc (str "Cannot infer an empty type for this expression") +| GTypArrow _ | GTypTuple _ -> + user_err ~loc (str "Type " ++ pr_glbtype t ++ str " is not an empty type") +| GTypRef (kn, _) -> + let def = Tac2env.interp_type kn in + match def with + | _, GTydAlg [] -> kn + | _ -> + user_err ~loc (str "Type " ++ pr_glbtype t ++ str " is not an empty type") + +let check_unit ?loc t = + let maybe_unit = match t with + | GTypVar _ -> true + | GTypArrow _ | GTypTuple _ -> false + | GTypRef (kn, _) -> KerName.equal kn t_unit + in + if not maybe_unit then warn_not_unit ?loc () + +let check_redundant_clause = function +| [] -> () +| (p, _) :: _ -> warn_redundant_clause ~loc:(loc_of_patexpr p) () + +let get_variable env (loc, qid) = + let (dp, id) = repr_qualid qid in + if DirPath.is_empty dp && Id.Map.mem id env.env_var then ArgVar (loc, id) + else + let kn = + try Tac2env.locate_ltac qid + with Not_found -> + CErrors.user_err ~loc (str "Unbound value " ++ pr_qualid qid) + in + ArgArg kn + +let get_constructor env (loc, qid) = + let c = try Some (Tac2env.locate_ltac qid) with Not_found -> None in + match c with + | Some knc -> + let kn = + try Tac2env.interp_constructor knc + with Not_found -> + CErrors.user_err ~loc (str "The term " ++ pr_qualid qid ++ + str " is not the constructor of an inductive type.") in + ArgArg (kn, knc) + | None -> + let (dp, id) = repr_qualid qid in + if DirPath.is_empty dp then ArgVar (loc, id) + else CErrors.user_err ~loc (str "Unbound constructor " ++ pr_qualid qid) + +let intern_atm env = function +| AtmInt n -> (GTacAtm (AtmInt n), GTypRef (t_int, [])) +| AtmStr s -> (GTacAtm (AtmStr s), GTypRef (t_string, [])) + +let invalid_pattern ~loc kn t = + let pt = match t with + | GCaseAlg kn' -> KerName.print kn + | GCaseTuple n -> str "tuple" + in + user_err ~loc (str "Invalid pattern, expected a pattern for type " ++ + KerName.print kn ++ str ", found a pattern of type " ++ pt) (** FIXME *) + +type pattern_kind = +| PKind_empty +| PKind_variant of KerName.t +| PKind_tuple of int +| PKind_any + +let get_pattern_kind env pl = match pl with +| [] -> PKind_empty +| p :: pl -> + let rec get_kind p pl = match fst p with + | CPatAny _ -> + begin match pl with + | [] -> PKind_any + | p :: pl -> get_kind p pl + end + | CPatRef (_, qid, []) -> + begin match get_constructor env qid with + | ArgVar _ -> + begin match pl with + | [] -> PKind_any + | p :: pl -> get_kind p pl + end + | ArgArg (data, _) -> PKind_variant data.cdata_type + end + | CPatRef (_, qid, _ :: _) -> + begin match get_constructor env qid with + | ArgVar (loc, _) -> + user_err ~loc (str "Unbound constructor " ++ pr_qualid (snd qid)) + | ArgArg (data, _) -> PKind_variant data.cdata_type + end + | CPatTup (_, tp) -> PKind_tuple (List.length tp) + in + get_kind p pl + +let rec intern_rec env = function +| CTacAtm (_, atm) -> intern_atm env atm +| CTacRef qid -> + begin match get_variable env qid with + | ArgVar (_, id) -> + let sch = Id.Map.find id env.env_var in + (GTacVar id, fresh_mix_type_scheme env sch) + | ArgArg kn -> + let (_, sch) = Tac2env.interp_global kn in + (GTacRef kn, fresh_type_scheme env sch) + end +| CTacFun (loc, bnd, e) -> + let fold (env, bnd, tl) ((_, na), t) = + let t = match t with + | None -> GTypVar (fresh_id env) + | Some t -> intern_type env t + in + let env = push_name na (monomorphic t) env in + (env, na :: bnd, t :: tl) + in + let (env, bnd, tl) = List.fold_left fold (env, [], []) bnd in + let (e, t) = intern_rec env e in + let t = List.fold_left (fun accu t -> GTypArrow (t, accu)) t tl in + (GTacFun (bnd, e), t) +| CTacApp (loc, f, args) -> + let (f, ft) = intern_rec env f in + let fold arg (args, t) = + let (arg, argt) = intern_rec env arg in + (arg :: args, GTypArrow (argt, t)) + in + let ret = GTypVar (fresh_id env) in + let (args, t) = List.fold_right fold args ([], ret) in + let () = unify loc env ft t in + (GTacApp (f, args), ret) +| CTacLet (loc, false, el, e) -> + let fold accu ((loc, na), _, _) = match na with + | Anonymous -> accu + | Name id -> + if Id.Set.mem id accu then + user_err ~loc (str "Variable " ++ Id.print id ++ str " is bound several \ + times in this matching") + else Id.Set.add id accu + in + let _ = List.fold_left fold Id.Set.empty el in + let fold ((loc, na), tc, e) (el, p) = + let (e, t) = intern_rec env e in + let () = match tc with + | None -> () + | Some tc -> + let tc = intern_type env tc in + unify loc env t tc + in + let t = if is_value e then abstract_var env t else monomorphic t in + ((na, e) :: el), ((na, t) :: p) + in + let (el, p) = List.fold_right fold el ([], []) in + let nenv = List.fold_left (fun accu (na, t) -> push_name na t env) env p in + let (e, t) = intern_rec nenv e in + (GTacLet (false, el, e), t) +| CTacLet (loc, true, el, e) -> + intern_let_rec env loc el e +| CTacTup (loc, []) -> + (GTacTup [], GTypRef (t_unit, [])) +| CTacTup (loc, el) -> + let fold e (el, tl) = + let (e, t) = intern_rec env e in + (e :: el, t :: tl) + in + let (el, tl) = List.fold_right fold el ([], []) in + (GTacTup el, GTypTuple tl) +| CTacArr (loc, []) -> + let id = fresh_id env in + (GTacArr [], GTypRef (t_int, [GTypVar id])) +| CTacArr (loc, e0 :: el) -> + let (e0, t0) = intern_rec env e0 in + let fold e el = + let loc = loc_of_tacexpr e in + let (e, t) = intern_rec env e in + let () = unify loc env t t0 in + e :: el + in + let el = e0 :: List.fold_right fold el [] in + (GTacArr el, GTypRef (t_array, [t0])) +| CTacLst (loc, []) -> + let id = fresh_id env in + (c_nil, GTypRef (t_list, [GTypVar id])) +| CTacLst (loc, e0 :: el) -> + let (e0, t0) = intern_rec env e0 in + let fold e el = + let loc = loc_of_tacexpr e in + let (e, t) = intern_rec env e in + let () = unify loc env t t0 in + c_cons e el + in + let el = c_cons e0 (List.fold_right fold el c_nil) in + (el, GTypRef (t_list, [t0])) +| CTacCnv (loc, e, tc) -> + let (e, t) = intern_rec env e in + let tc = intern_type env tc in + let () = unify loc env t tc in + (e, tc) +| CTacSeq (loc, e1, e2) -> + let (e1, t1) = intern_rec env e1 in + let (e2, t2) = intern_rec env e2 in + let () = check_elt_unit loc env t1 in + (GTacLet (false, [Anonymous, e1], e2), t2) +| CTacCse (loc, e, pl) -> + intern_case env loc e pl + +and intern_let_rec env loc el e = + let fold accu ((loc, na), _, _) = match na with + | Anonymous -> accu + | Name id -> + if Id.Set.mem id accu then + user_err ~loc (str "Variable " ++ Id.print id ++ str " is bound several \ + times in this matching") + else Id.Set.add id accu + in + let _ = List.fold_left fold Id.Set.empty el in + let map env ((loc, na), t, e) = + let id = fresh_id env in + let env = push_name na (monomorphic (GTypVar id)) env in + (env, (loc, na, t, e, id)) + in + let (env, el) = List.fold_map map env el in + let fold (loc, na, tc, e, id) (el, tl) = + let loc_e = loc_of_tacexpr e in + let (e, t) = intern_rec env e in + let () = + if not (is_rec_rhs e) then + user_err ~loc:loc_e (str "This kind of expression is not allowed as \ + right-hand side of a recursive binding") + in + let () = unify loc env t (GTypVar id) in + let () = match tc with + | None -> () + | Some tc -> + let tc = intern_type env tc in + unify loc env t tc + in + ((na, e) :: el, t :: tl) + in + let (el, tl) = List.fold_right fold el ([], []) in + let (e, t) = intern_rec env e in + (GTacLet (true, el, e), t) + +(** For now, patterns recognized by the pattern-matching compiling are limited + to depth-one where leaves are either variables or catch-all *) +and intern_case env loc e pl = + let (e', t) = intern_rec env e in + let todo ~loc () = user_err ~loc (str "Pattern not handled yet") in + match get_pattern_kind env pl with + | PKind_any -> + let (pat, b) = List.hd pl in + let na = match pat with + | CPatAny _ -> Anonymous + | CPatRef (_, (_, qid), _) -> Name (snd (repr_qualid qid)) + | _ -> assert false + in + let () = check_redundant_clause (List.tl pl) in + let env = push_name na (monomorphic t) env in + let (b, tb) = intern_rec env b in + (GTacLet (false, [na, e'], b), tb) + | PKind_empty -> + let kn = check_elt_empty loc env t in + let r = fresh_id env in + (GTacCse (e', GCaseAlg kn, [||], [||]), GTypVar r) + | PKind_tuple len -> + begin match pl with + | [] -> assert false + | [CPatTup (_, []), b] -> + let () = unify (loc_of_tacexpr e) env t (GTypRef (t_unit, [])) in + let (b, tb) = intern_rec env b in + (GTacCse (e', GCaseAlg t_unit, [|b|], [||]), tb) + | [CPatTup (_, pl), b] -> + let map = function + | CPatAny _ -> Anonymous + | CPatRef (loc, qid, []) -> + begin match get_constructor env qid with + | ArgVar (_, id) -> Name id + | ArgArg _ -> todo ~loc () + end + | p -> todo ~loc:(loc_of_patexpr p) () + in + let ids = Array.map_of_list map pl in + let tc = GTypTuple (List.map (fun _ -> GTypVar (fresh_id env)) pl) in + let () = unify (loc_of_tacexpr e) env t tc in + let (b, tb) = intern_rec env b in + (GTacCse (e', GCaseTuple len, [||], [|ids, b|]), tb) + | (p, _) :: _ -> todo ~loc:(loc_of_patexpr p) () + end + | PKind_variant kn -> + let subst, tc = fresh_reftype env kn in + let () = unify (loc_of_tacexpr e) env t tc in + let (params, def) = Tac2env.interp_type kn in + let cstrs = match def with + | GTydAlg c -> c + | _ -> assert false + in + let count (const, nonconst) (c, args) = match args with + | [] -> (succ const, nonconst) + | _ :: _ -> (const, succ nonconst) + in + let nconst, nnonconst = List.fold_left count (0, 0) cstrs in + let const = Array.make nconst None in + let nonconst = Array.make nnonconst None in + let ret = GTypVar (fresh_id env) in + let rec intern_branch = function + | [] -> () + | (pat, br) :: rem -> + let tbr = match pat with + | CPatAny _ -> + let () = check_redundant_clause rem in + let (br', brT) = intern_rec env br in + (** Fill all remaining branches *) + let fill (ncst, narg) (_, args) = + if List.is_empty args then + let () = + if Option.is_empty const.(ncst) then const.(ncst) <- Some br' + in + (succ ncst, narg) + else + let () = + if Option.is_empty const.(narg) then + let ids = Array.map_of_list (fun _ -> Anonymous) args in + nonconst.(narg) <- Some (ids, br') + in + (ncst, succ narg) + in + let _ = List.fold_left fill (0, 0) cstrs in + brT + | CPatRef (loc, qid, args) -> + let data = match get_constructor env qid with + | ArgVar _ -> todo ~loc () + | ArgArg (data, _) -> + let () = + let kn' = data.cdata_type in + if not (KerName.equal kn kn') then + invalid_pattern ~loc kn (GCaseAlg kn') + in + data + in + let get_id = function + | CPatAny _ -> Anonymous + | CPatRef (loc, qid, []) -> + begin match get_constructor env qid with + | ArgVar (_, id) -> Name id + | ArgArg _ -> todo ~loc () + end + | p -> todo ~loc:(loc_of_patexpr p) () + in + let ids = List.map get_id args in + let nids = List.length ids in + let nargs = List.length (snd data.cdata_args) in + let () = + if not (Int.equal nids nargs) then + user_err ~loc (str "Constructor expects " ++ int nargs ++ + str " arguments, but is applied to " ++ int nids ++ + str " arguments") + in + let fold env id tpe = + (** Instantiate all arguments *) + let subst n = GTypVar subst.(n) in + let tpe = subst_type subst tpe in + push_name id (monomorphic tpe) env + in + let nenv = List.fold_left2 fold env ids (snd data.cdata_args) in + let (br', brT) = intern_rec nenv br in + let () = + let index = data.cdata_indx in + if List.is_empty args then + if Option.is_empty const.(index) then const.(index) <- Some br' + else warn_redundant_clause ~loc () + else + let ids = Array.of_list ids in + if Option.is_empty nonconst.(index) then nonconst.(index) <- Some (ids, br') + else warn_redundant_clause ~loc () + in + brT + | CPatTup (loc, tup) -> + invalid_pattern ~loc kn (GCaseTuple (List.length tup)) + in + let () = unify (loc_of_tacexpr br) env ret tbr in + intern_branch rem + in + let () = intern_branch pl in + let map = function + | None -> user_err ~loc (str "Unhandled match case") (** FIXME *) + | Some x -> x + in + let const = Array.map map const in + let nonconst = Array.map map nonconst in + let ce = GTacCse (e', GCaseAlg kn, const, nonconst) in + (ce, ret) + +let normalize env (count, vars) (t : UF.elt glb_typexpr) = + let get_var id = + try UF.Map.find id !vars + with Not_found -> + let () = assert env.env_opn in + let n = GTypVar !count in + let () = incr count in + let () = vars := UF.Map.add id n !vars in + n + in + let rec subst id = match UF.find id env.env_cst with + | id, None -> get_var id + | _, Some t -> subst_type subst t + in + subst_type subst t + +let intern e = + let env = empty_env () in + let (e, t) = intern_rec env e in + let count = ref 0 in + let vars = ref UF.Map.empty in + let t = normalize env (count, vars) t in + (e, (!count, t)) + +let intern_typedef self (ids, t) : glb_quant_typedef = + let env = { (empty_env ()) with env_rec = self } in + (** Initialize type parameters *) + let map id = get_alias id env in + let ids = List.map map ids in + let count = ref (List.length ids) in + let vars = ref UF.Map.empty in + let iter n id = vars := UF.Map.add id (GTypVar n) !vars in + let () = List.iteri iter ids in + (** Do not accept unbound type variables *) + let env = { env with env_opn = false } in + let intern t = + let t = intern_type env t in + normalize env (count, vars) t + in + let count = !count in + match t with + | CTydDef None -> (count, GTydDef None) + | CTydDef (Some t) -> (count, GTydDef (Some (intern t))) + | CTydAlg constrs -> + let map (c, t) = (c, List.map intern t) in + let constrs = List.map map constrs in + (count, GTydAlg constrs) + | CTydRec fields -> + let map (c, mut, t) = (c, mut, intern t) in + let fields = List.map map fields in + (count, GTydRec fields) + +let intern_open_type t = + let env = empty_env () in + let t = intern_type env t in + let count = ref 0 in + let vars = ref UF.Map.empty in + let t = normalize env (count, vars) t in + (!count, t) + +(** Kernel substitution *) + +open Mod_subst + +let rec subst_type subst t = match t with +| GTypVar _ -> t +| GTypArrow (t1, t2) -> + let t1' = subst_type subst t1 in + let t2' = subst_type subst t2 in + if t1' == t1 && t2' == t2 then t + else GTypArrow (t1', t2') +| GTypTuple tl -> + let tl'= List.smartmap (fun t -> subst_type subst t) tl in + if tl' == tl then t else GTypTuple tl' +| GTypRef (kn, tl) -> + let kn' = subst_kn subst kn in + let tl' = List.smartmap (fun t -> subst_type subst t) tl in + if kn' == kn && tl' == tl then t else GTypRef (kn', tl') + +let subst_case_info subst ci = match ci with +| GCaseAlg kn -> + let kn' = subst_kn subst kn in + if kn' == kn then ci else GCaseAlg kn' +| GCaseTuple _ -> ci + +let rec subst_expr subst e = match e with +| GTacAtm _ | GTacVar _ | GTacPrm _ -> e +| GTacRef kn -> GTacRef (subst_kn subst kn) +| GTacFun (ids, e) -> GTacFun (ids, subst_expr subst e) +| GTacApp (f, args) -> + GTacApp (subst_expr subst f, List.map (fun e -> subst_expr subst e) args) +| GTacLet (r, bs, e) -> + let bs = List.map (fun (na, e) -> (na, subst_expr subst e)) bs in + GTacLet (r, bs, subst_expr subst e) +| GTacTup el -> + GTacTup (List.map (fun e -> subst_expr subst e) el) +| GTacArr el -> + GTacArr (List.map (fun e -> subst_expr subst e) el) +| GTacCst (kn, n, el) -> + GTacCst (subst_kn subst kn, n, List.map (fun e -> subst_expr subst e) el) +| GTacCse (e, ci, cse0, cse1) -> + let cse0' = Array.map (fun e -> subst_expr subst e) cse0 in + let cse1' = Array.map (fun (ids, e) -> (ids, subst_expr subst e)) cse1 in + let ci' = subst_case_info subst ci in + GTacCse (subst_expr subst e, ci', cse0', cse1') + +let subst_typedef subst e = match e with +| GTydDef t -> + let t' = Option.smartmap (fun t -> subst_type subst t) t in + if t' == t then e else GTydDef t' +| GTydAlg constrs -> + let map (c, tl as p) = + let tl' = List.smartmap (fun t -> subst_type subst t) tl in + if tl' == tl then p else (c, tl') + in + let constrs' = List.smartmap map constrs in + if constrs' == constrs then e else GTydAlg constrs' +| GTydRec fields -> + let map (c, mut, t as p) = + let t' = subst_type subst t in + if t' == t then p else (c, mut, t') + in + let fields' = List.smartmap map fields in + if fields' == fields then e else GTydRec fields' + +let subst_quant_typedef subst (prm, def as qdef) = + let def' = subst_typedef subst def in + if def' == def then qdef else (prm, def') + +let subst_type_scheme subst (prm, t as sch) = + let t' = subst_type subst t in + if t' == t then sch else (prm, t') diff --git a/tac2intern.mli b/tac2intern.mli new file mode 100644 index 0000000000..a6be01d647 --- /dev/null +++ b/tac2intern.mli @@ -0,0 +1,30 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Loc.t + +val intern : raw_tacexpr -> glb_tacexpr * type_scheme +val intern_typedef : (KerName.t * int) Id.Map.t -> raw_quant_typedef -> glb_quant_typedef +val intern_open_type : raw_typexpr -> type_scheme + +(** Check that a term is a value. Only values are safe to marshall between + processes. *) +val is_value : glb_tacexpr -> bool +val check_unit : ?loc:Loc.t -> int glb_typexpr -> unit + +val subst_type : substitution -> 'a glb_typexpr -> 'a glb_typexpr +val subst_expr : substitution -> glb_tacexpr -> glb_tacexpr +val subst_quant_typedef : substitution -> glb_quant_typedef -> glb_quant_typedef +val subst_type_scheme : substitution -> type_scheme -> type_scheme + +val pr_glbtype : 'a glb_typexpr -> Pp.std_ppcmds diff --git a/tac2interp.ml b/tac2interp.ml new file mode 100644 index 0000000000..221f107dc8 --- /dev/null +++ b/tac2interp.ml @@ -0,0 +1,108 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* ist +| Name id -> Id.Map.add id v ist + +let get_var ist id = + try Id.Map.find id ist with Not_found -> + anomaly (str "Unbound variable " ++ Id.print id) + +let get_ref ist kn = + try fst (Tac2env.interp_global kn) with Not_found -> + anomaly (str "Unbound reference" ++ KerName.print kn) + +let return = Proofview.tclUNIT + +let rec interp ist = function +| GTacAtm (AtmInt n) -> return (ValInt n) +| GTacAtm (AtmStr s) -> return (ValStr s) +| GTacVar id -> return (get_var ist id) +| GTacRef qid -> return (get_ref ist qid) +| GTacFun (ids, e) -> + let cls = { clos_env = ist; clos_var = ids; clos_exp = e } in + return (ValCls cls) +| GTacApp (f, args) -> + interp ist f >>= fun f -> + Proofview.Monad.List.map (fun e -> interp ist e) args >>= fun args -> + interp_app ist f args +| GTacLet (false, el, e) -> + let fold accu (na, e) = + interp ist e >>= fun e -> + return (push_name accu na e) + in + Proofview.Monad.List.fold_left fold ist el >>= fun ist -> + interp ist e +| GTacLet (true, el, e) -> + let map (na, e) = match e with + | GTacFun (ids, e) -> + let cls = { clos_env = ist; clos_var = ids; clos_exp = e } in + na, cls + | _ -> anomaly (str "Ill-formed recursive function") + in + let fixs = List.map map el in + let fold accu (na, cls) = match na with + | Anonymous -> accu + | Name id -> Id.Map.add id (ValCls cls) accu + in + let ist = List.fold_left fold ist fixs in + (** Hack to make a cycle imperatively in the environment *) + let iter (_, e) = e.clos_env <- ist in + let () = List.iter iter fixs in + interp ist e +| GTacTup [] -> + return (ValInt 0) +| GTacTup el | GTacArr el -> + Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> + return (ValBlk (0, Array.of_list el)) +| GTacCst (_, n, []) -> return (ValInt n) +| GTacCst (_, n, el) -> + Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> + return (ValBlk (n, Array.of_list el)) +| GTacCse (e, _, cse0, cse1) -> + interp ist e >>= fun e -> interp_case ist e cse0 cse1 +| GTacPrm (ml, el) -> + Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> + Tac2env.interp_primitive ml el + +and interp_app ist f args = match f with +| ValCls { clos_env = ist; clos_var = ids; clos_exp = e } -> + let rec push ist ids args = match ids, args with + | [], [] -> interp ist e + | [], _ :: _ -> interp ist e >>= fun f -> interp_app ist f args + | _ :: _, [] -> + let cls = { clos_env = ist; clos_var = ids; clos_exp = e } in + return (ValCls cls) + | id :: ids, arg :: args -> push (push_name ist id arg) ids args + in + push ist ids args +| ValExt _ | ValInt _ | ValBlk _ | ValStr _ -> + anomaly (str "Unexpected value shape") + +and interp_case ist e cse0 cse1 = match e with +| ValInt n -> interp ist cse0.(n) +| ValBlk (n, args) -> + let (ids, e) = cse1.(n) in + let ist = CArray.fold_left2 push_name ist ids args in + interp ist e +| ValExt _ | ValStr _ | ValCls _ -> + anomaly (str "Unexpected value shape") diff --git a/tac2interp.mli b/tac2interp.mli new file mode 100644 index 0000000000..b11ee36012 --- /dev/null +++ b/tac2interp.mli @@ -0,0 +1,19 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* glb_tacexpr -> valexpr Proofview.tactic diff --git a/vo.itarget b/vo.itarget new file mode 100644 index 0000000000..776404ad79 --- /dev/null +++ b/vo.itarget @@ -0,0 +1 @@ +Ltac2.vo -- cgit v1.2.3 From 0c3c2459eae24cc8e87c7c6a4a4e6a1afd171d72 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 25 Oct 2016 15:11:41 +0200 Subject: Embedding generic arguments in Ltac2 AST. --- tac2env.ml | 19 ++++++++++++++++++- tac2env.mli | 10 ++++++++++ tac2expr.mli | 4 ++++ tac2intern.ml | 22 ++++++++++++++++++++-- tac2interp.ml | 4 ++++ 5 files changed, 56 insertions(+), 3 deletions(-) diff --git a/tac2env.ml b/tac2env.ml index d17e657516..3500b2ef3d 100644 --- a/tac2env.ml +++ b/tac2env.ml @@ -52,7 +52,7 @@ let rec eval_pure = function | GTacCst (_, n, []) -> ValInt n | GTacCst (_, n, el) -> ValBlk (n, Array.map_of_list eval_pure el) | GTacAtm (AtmStr _) | GTacArr _ | GTacLet _ | GTacVar _ -| GTacApp _ | GTacCse _ | GTacPrm _ -> +| GTacApp _ | GTacCse _ | GTacPrm _ | GTacExt _ -> anomaly (Pp.str "Term is not a syntactical value") let define_global kn e = @@ -137,3 +137,20 @@ let locate_type qid = let locate_extended_all_type qid = let tab = !nametab in KnTab.find_prefixes qid tab.tab_type + +type 'a ml_object = { + ml_type : type_constant; + ml_interp : environment -> 'a -> Geninterp.Val.t Proofview.tactic; +} + +module MLTypeObj = +struct + type ('a, 'b, 'c) obj = 'b ml_object + let name = "ltac2_ml_type" + let default _ = None +end + +module MLType = Genarg.Register(MLTypeObj) + +let define_ml_object t tpe = MLType.register0 t tpe +let interp_ml_object t = MLType.obj t diff --git a/tac2env.mli b/tac2env.mli index efabdb5466..eb471b9abf 100644 --- a/tac2env.mli +++ b/tac2env.mli @@ -56,3 +56,13 @@ val locate_extended_all_type : qualid -> type_constant list val define_primitive : ml_tactic_name -> ml_tactic -> unit val interp_primitive : ml_tactic_name -> ml_tactic + +(** {5 ML primitive types} *) + +type 'a ml_object = { + ml_type : type_constant; + ml_interp : environment -> 'a -> Geninterp.Val.t Proofview.tactic; +} + +val define_ml_object : ('a, 'b, 'c) genarg_type -> 'b ml_object -> unit +val interp_ml_object : ('a, 'b, 'c) genarg_type -> 'b ml_object diff --git a/tac2expr.mli b/tac2expr.mli index 445c69aa23..15c630ca87 100644 --- a/tac2expr.mli +++ b/tac2expr.mli @@ -76,6 +76,7 @@ type raw_tacexpr = | CTacCnv of Loc.t * raw_tacexpr * raw_typexpr | CTacSeq of Loc.t * raw_tacexpr * raw_tacexpr | CTacCse of Loc.t * raw_tacexpr * raw_taccase list +| CTacExt of Loc.t * raw_generic_argument and raw_taccase = raw_patexpr * raw_tacexpr @@ -94,6 +95,7 @@ type glb_tacexpr = | GTacArr of glb_tacexpr list | GTacCst of KerName.t * int * glb_tacexpr list | GTacCse of glb_tacexpr * case_info * glb_tacexpr array * (Name.t array * glb_tacexpr) array +| GTacExt of glob_generic_argument | GTacPrm of ml_tactic_name * glb_tacexpr list (** Toplevel statements *) @@ -134,3 +136,5 @@ and closure = { } type ml_tactic = valexpr list -> valexpr Proofview.tactic + +type environment = valexpr Id.Map.t diff --git a/tac2intern.ml b/tac2intern.ml index b1b35b0787..516690dfd3 100644 --- a/tac2intern.ml +++ b/tac2intern.ml @@ -144,6 +144,9 @@ let empty_env () = { env_rec = Id.Map.empty; } +let ltac2_env : environment Genintern.Store.field = + Genintern.Store.field () + let fresh_id env = UF.fresh env.env_cst let get_alias (loc, id) env = @@ -171,6 +174,7 @@ let loc_of_tacexpr = function | CTacCnv (loc, _, _) -> loc | CTacSeq (loc, _, _) -> loc | CTacCse (loc, _, _) -> loc +| CTacExt (loc, _) -> loc let loc_of_patexpr = function | CPatAny loc -> loc @@ -324,12 +328,12 @@ let rec is_value = function | GTacCst (kn, n, el) -> (** To be a value, a constructor must be immutable *) assert false (** TODO *) -| GTacArr _ | GTacCse _ | GTacPrm _ -> false +| GTacArr _ | GTacCse _ | GTacExt _ | GTacPrm _ -> false let is_rec_rhs = function | GTacFun _ -> true | GTacAtm _ | GTacVar _ | GTacRef _ | GTacApp _ | GTacLet _ -> false -| GTacTup _ | GTacArr _ | GTacPrm _ | GTacCst _ | GTacCse _ -> false +| GTacTup _ | GTacArr _ | GTacExt _ | GTacPrm _ | GTacCst _ | GTacCse _ -> false let rec fv_type f t accu = match t with | GTypVar id -> f id accu @@ -600,6 +604,17 @@ let rec intern_rec env = function (GTacLet (false, [Anonymous, e1], e2), t2) | CTacCse (loc, e, pl) -> intern_case env loc e pl +| CTacExt (loc, ext) -> + let open Genintern in + let GenArg (Rawwit tag, _) = ext in + let tpe = interp_ml_object tag in + (** External objects do not have access to the named context because this is + not stable by dynamic semantics. *) + let genv = Global.env_of_context Environ.empty_named_context_val in + let ist = empty_glob_sign genv in + let ist = { ist with extra = Store.set ist.extra ltac2_env env } in + let (_, ext) = generic_intern ist ext in + (GTacExt ext, GTypRef (tpe.ml_type, [])) and intern_let_rec env loc el e = let fold accu ((loc, na), _, _) = match na with @@ -892,6 +907,9 @@ let rec subst_expr subst e = match e with let cse1' = Array.map (fun (ids, e) -> (ids, subst_expr subst e)) cse1 in let ci' = subst_case_info subst ci in GTacCse (subst_expr subst e, ci', cse0', cse1') +| GTacExt ext -> + let ext' = Genintern.generic_substitute subst ext in + if ext' == ext then e else GTacExt ext' let subst_typedef subst e = match e with | GTydDef t -> diff --git a/tac2interp.ml b/tac2interp.ml index 221f107dc8..b868caf963 100644 --- a/tac2interp.ml +++ b/tac2interp.ml @@ -83,6 +83,10 @@ let rec interp ist = function | GTacPrm (ml, el) -> Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> Tac2env.interp_primitive ml el +| GTacExt e -> + let GenArg (Glbwit tag, e) = e in + let tpe = Tac2env.interp_ml_object tag in + tpe.Tac2env.ml_interp ist e >>= fun e -> return (ValExt e) and interp_app ist f args = match f with | ValCls { clos_env = ist; clos_var = ids; clos_exp = e } -> -- cgit v1.2.3 From 2dc3175916f3968d4cdba9af140fbc2667ff70a5 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 2 Dec 2016 14:31:27 +0100 Subject: Allowing to include Coq terms in Ltac2 using the constr:(...) syntax. --- g_ltac2.ml4 | 4 ++++ tac2core.ml | 69 +++++++++++++++++++++++++++++++++++++++++++++++++++++++++++ tac2intern.ml | 2 +- 3 files changed, 74 insertions(+), 1 deletion(-) diff --git a/g_ltac2.ml4 b/g_ltac2.ml4 index 349220f9de..b613f22a8d 100644 --- a/g_ltac2.ml4 +++ b/g_ltac2.ml4 @@ -21,6 +21,9 @@ let tac2def_typ = Gram.entry_create "tactic:tac2def_typ" let tac2def_ext = Gram.entry_create "tactic:tac2def_ext" let tac2mode = Gram.entry_create "vernac:ltac2_command" +let inj_wit wit loc x = CTacExt (loc, Genarg.in_gen (Genarg.rawwit wit) x) +let inj_constr loc c = inj_wit Stdarg.wit_constr loc c + GEXTEND Gram GLOBAL: tac2expr tac2type tac2def_val tac2def_typ tac2def_ext; tac2pat: @@ -77,6 +80,7 @@ GEXTEND Gram [ [ n = Prim.integer -> CTacAtm (!@loc, AtmInt n) | s = Prim.string -> CTacAtm (!@loc, AtmStr s) | id = Prim.qualid -> CTacRef id + | IDENT "constr"; ":"; "("; c = Constr.lconstr; ")" -> inj_constr !@loc c ] ] ; let_clause: diff --git a/tac2core.ml b/tac2core.ml index af4124e647..95632bf7b1 100644 --- a/tac2core.ml +++ b/tac2core.ml @@ -8,8 +8,11 @@ open CSig open Pp +open CErrors open Names +open Genarg open Geninterp +open Tac2env open Tac2expr open Proofview.Notations @@ -26,6 +29,7 @@ let t_string = coq_core "string" let t_array = coq_core "array" let t_unit = coq_core "unit" let t_list = coq_core "list" +let t_constr = coq_core "constr" let c_nil = coq_core "[]" let c_cons = coq_core "::" @@ -66,6 +70,8 @@ match Val.eq tag tag' with let val_pp = Val.create "ltac2:pp" let get_pp v = extract_val val_pp v +let val_valexpr = Val.create "ltac2:valexpr" + (** Helper functions *) let return x = Proofview.tclUNIT x @@ -117,3 +123,66 @@ let () = Tac2env.define_primitive (pname "array_make") prm_array_make let () = Tac2env.define_primitive (pname "array_length") prm_array_length let () = Tac2env.define_primitive (pname "array_get") prm_array_get let () = Tac2env.define_primitive (pname "array_set") prm_array_set + +(** ML types *) + +let val_tag t = match val_tag t with +| Val.Base t -> t +| _ -> assert false + +let tag_constr = val_tag (topwit Stdarg.wit_constr) + +let constr_flags () = + let open Pretyping in + { + use_typeclasses = true; + solve_unification_constraints = true; + use_hook = Pfedit.solve_by_implicit_tactic (); + fail_evar = true; + expand_evars = true + } + +(** In Ltac2, the notion of "current environment" only makes sense when there is + at most one goal under focus. Contrarily to Ltac1, instead of dynamically + focussing when we need it, we raise a non-backtracking error when it does + not make sense. *) +exception NonFocussedGoal + +let () = register_handler begin function +| NonFocussedGoal -> str "Several goals under focus" +| _ -> raise Unhandled +end + +let pf_apply f = + Proofview.Goal.goals >>= function + | [] -> + Proofview.tclENV >>= fun env -> + Proofview.tclEVARMAP >>= fun sigma -> + f env sigma + | [gl] -> + gl >>= fun gl -> + f (Proofview.Goal.env gl) (Tacmach.New.project gl) + | _ :: _ :: _ -> + Proofview.tclLIFT (Proofview.NonLogical.raise NonFocussedGoal) + +(** Embed all Ltac2 data into Values *) +let to_lvar ist = + let open Pretyping in + let map e = Val.Dyn (val_valexpr, e) in + let lfun = Id.Map.map map ist in + { empty_lvar with ltac_genargs = lfun } + +let () = + let open Pretyping in + let interp ist (c, _) = pf_apply begin fun env sigma -> + let ist = to_lvar ist in + let (sigma, c) = understand_ltac (constr_flags ()) env sigma ist WithoutTypeConstraint c in + let c = Val.Dyn (tag_constr, c) in + Proofview.Unsafe.tclEVARS sigma >>= fun () -> + Proofview.tclUNIT c + end in + let obj = { + ml_type = t_constr; + ml_interp = interp; + } in + define_ml_object Stdarg.wit_constr obj diff --git a/tac2intern.ml b/tac2intern.ml index 516690dfd3..a554f959a0 100644 --- a/tac2intern.ml +++ b/tac2intern.ml @@ -613,7 +613,7 @@ let rec intern_rec env = function let genv = Global.env_of_context Environ.empty_named_context_val in let ist = empty_glob_sign genv in let ist = { ist with extra = Store.set ist.extra ltac2_env env } in - let (_, ext) = generic_intern ist ext in + let (_, ext) = Flags.with_option Ltac_plugin.Tacintern.strict_check (fun () -> generic_intern ist ext) () in (GTacExt ext, GTypRef (tpe.ml_type, [])) and intern_let_rec env loc el e = -- cgit v1.2.3 From 07ad1ca45473ba02db9b687bb7e89d440ba79b20 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 4 Dec 2016 21:48:06 +0100 Subject: Proper handling of record types. We add the standard ML facilities for records, that is, projections, mutable fields and primitive to set them. --- Ltac2.v | 2 +- g_ltac2.ml4 | 30 ++++++++-- tac2entries.ml | 79 +++++++++++++++++-------- tac2env.ml | 65 ++++++++++++++++++--- tac2env.mli | 45 ++++++++++++--- tac2expr.mli | 20 +++++-- tac2intern.ml | 180 +++++++++++++++++++++++++++++++++++++++++++++++---------- tac2interp.ml | 19 ++++++ 8 files changed, 360 insertions(+), 80 deletions(-) diff --git a/Ltac2.v b/Ltac2.v index a952524e71..0933c1e0b4 100644 --- a/Ltac2.v +++ b/Ltac2.v @@ -20,7 +20,7 @@ Ltac2 Type 'a array. (** Pervasive types *) -Ltac2 Type 'a option := | None | Some ('a). +Ltac2 Type 'a option := [ None | Some ('a) ]. (** Primitive tactics *) diff --git a/g_ltac2.ml4 b/g_ltac2.ml4 index b613f22a8d..ce2becd9f9 100644 --- a/g_ltac2.ml4 +++ b/g_ltac2.ml4 @@ -50,6 +50,8 @@ GEXTEND Gram [ e1 = tac2expr; ";"; e2 = tac2expr -> CTacSeq (!@loc, e1, e2) ] | "1" LEFTA [ e = tac2expr; el = LIST1 tac2expr LEVEL "0" -> CTacApp (!@loc, e, el) + | e = SELF; ".("; qid = Prim.qualid; ")" -> CTacPrj (!@loc, e, qid) + | e = SELF; ".("; qid = Prim.qualid; ")"; ":="; r = tac2expr LEVEL "1" -> CTacSet (!@loc, e, qid, r) | e0 = tac2expr; ","; el = LIST1 tac2expr LEVEL "0" SEP "," -> CTacTup (!@loc, e0 :: el) ] | "0" [ "("; a = tac2expr LEVEL "5"; ")" -> a @@ -57,6 +59,7 @@ GEXTEND Gram | "()" -> CTacTup (!@loc, []) | "("; ")" -> CTacTup (!@loc, []) | "["; a = LIST0 tac2expr LEVEL "1" SEP ";"; "]" -> CTacLst (!@loc, a) + | "{"; a = tac2rec_fieldexprs; "}" -> CTacRec (!@loc, a) | a = tactic_atom -> a ] ] ; @@ -126,16 +129,35 @@ GEXTEND Gram ; tac2typ_knd: [ [ t = tac2type -> CTydDef (Some t) - | t = tac2alg_type -> CTydAlg t ] ] + | "["; t = tac2alg_constructors; "]" -> CTydAlg t + | "{"; t = tac2rec_fields; "}"-> CTydRec t ] ] ; - tac2alg_type: - [ [ -> [] - | "|"; bl = LIST1 tac2alg_constructor SEP "|" -> bl ] ] + tac2alg_constructors: + [ [ "|"; cs = LIST1 tac2alg_constructor SEP "|" -> cs + | cs = LIST0 tac2alg_constructor SEP "|" -> cs ] ] ; tac2alg_constructor: [ [ c = Prim.ident -> (c, []) | c = Prim.ident; "("; args = LIST0 tac2type SEP ","; ")"-> (c, args) ] ] ; + tac2rec_fields: + [ [ f = tac2rec_field; ";"; l = tac2rec_fields -> f :: l + | f = tac2rec_field; ";" -> [f] + | f = tac2rec_field -> [f] + | -> [] ] ] + ; + tac2rec_field: + [ [ mut = [ -> false | IDENT "mutable" -> true]; id = Prim.ident; ":"; t = tac2type -> (id, mut, t) ] ] + ; + tac2rec_fieldexprs: + [ [ f = tac2rec_fieldexpr; ";"; l = tac2rec_fieldexprs -> f :: l + | f = tac2rec_fieldexpr; ";" -> [f] + | f = tac2rec_fieldexpr-> [f] + | -> [] ] ] + ; + tac2rec_fieldexpr: + [ [ qid = Prim.qualid; ":="; e = tac2expr LEVEL "1" -> qid, e ] ] + ; tac2typ_prm: [ [ -> [] | id = typ_param -> [!@loc, id] diff --git a/tac2entries.ml b/tac2entries.ml index 7572270ab3..4098324f12 100644 --- a/tac2entries.ml +++ b/tac2entries.ml @@ -13,6 +13,7 @@ open Names open Libnames open Libobject open Nametab +open Tac2env open Tac2expr open Tac2intern open Vernacexpr @@ -24,14 +25,14 @@ type tacdef = { } let perform_tacdef visibility ((sp, kn), def) = - let () = if not def.tacdef_local then Tac2env.push_ltac visibility sp kn in + let () = if not def.tacdef_local then Tac2env.push_ltac visibility sp (TacConstant kn) in Tac2env.define_global kn (def.tacdef_expr, def.tacdef_type) let load_tacdef i obj = perform_tacdef (Until i) obj let open_tacdef i obj = perform_tacdef (Exactly i) obj let cache_tacdef ((sp, kn), def) = - let () = Tac2env.push_ltac (Until 1) sp kn in + let () = Tac2env.push_ltac (Until 1) sp (TacConstant kn) in Tac2env.define_global kn (def.tacdef_expr, def.tacdef_type) let subst_tacdef (subst, def) = @@ -71,12 +72,19 @@ let push_typedef visibility sp kn (_, def) = match def with let iter (c, _) = let spc = change_sp_label sp c in let knc = change_kn_label kn c in - Tac2env.push_ltac visibility spc knc + Tac2env.push_ltac visibility spc (TacConstructor knc) in Tac2env.push_type visibility sp kn; List.iter iter cstrs -| GTydRec _ -> - assert false (** FIXME *) +| GTydRec fields -> + (** Register fields *) + let iter (c, _, _) = + let spc = change_sp_label sp c in + let knc = change_kn_label kn c in + Tac2env.push_projection visibility spc knc + in + Tac2env.push_type visibility sp kn; + List.iter iter fields let next i = let ans = !i in @@ -95,27 +103,31 @@ let define_typedef kn (params, def as qdef) = match def with let iter (c, args) = let knc = change_kn_label kn c in let tag = if List.is_empty args then next constant else next nonconstant in - let ids = List.mapi (fun i _ -> dummy_var i) args in - let c = GTacCst (kn, tag, List.map (fun id -> GTacVar id) ids) in - let c = - if List.is_empty args then c - else GTacFun (List.map (fun id -> Name id) ids, c) - in - let fold arg tpe = GTypArrow (arg, tpe) in - let cT = GTypRef (kn, List.init params (fun i -> GTypVar i)) in - let cT = List.fold_right fold args cT in let data = { - Tac2env.cdata_type = kn; - cdata_args = params, args; + Tac2env.cdata_prms = params; + cdata_type = kn; + cdata_args = args; cdata_indx = tag; } in - Tac2env.define_constructor knc data; - Tac2env.define_global knc (c, (params, cT)) + Tac2env.define_constructor knc data in Tac2env.define_type kn qdef; List.iter iter cstrs -| GTydRec _ -> - assert false (** FIXME *) +| GTydRec fs -> + (** Define projections *) + let iter i (id, mut, t) = + let knp = change_kn_label kn id in + let proj = { + Tac2env.pdata_prms = params; + pdata_type = kn; + pdata_ptyp = t; + pdata_mutb = mut; + pdata_indx = i; + } in + Tac2env.define_projection knp proj + in + Tac2env.define_type kn qdef; + List.iteri iter fs let perform_typdef vs ((sp, kn), def) = let () = if not def.typdef_local then push_typedef vs sp kn def.typdef_expr in @@ -213,8 +225,22 @@ let register_type ?(local = false) isrec types = if isrec then user_err ~loc (str "The type abbreviation " ++ Id.print id ++ str " cannot be recursive") - | CTydAlg _ -> () (** FIXME *) - | CTydRec _ -> assert false (** FIXME *) + | CTydAlg cs -> + let same_name (id1, _) (id2, _) = Id.equal id1 id2 in + let () = match List.duplicates same_name cs with + | [] -> () + | (id, _) :: _ -> + user_err (str "Multiple definitions of the constructor " ++ Id.print id) + in + () + | CTydRec ps -> + let same_name (id1, _, _) (id2, _, _) = Id.equal id1 id2 in + let () = match List.duplicates same_name ps with + | [] -> () + | (id, _, _) :: _ -> + user_err (str "Multiple definitions of the projection " ++ Id.print id) + in + () in let () = List.iter check types in let self = @@ -275,8 +301,13 @@ let print_ltac ref = try Tac2env.locate_ltac qid with Not_found -> user_err ~loc (str "Unknown tactic " ++ pr_qualid qid) in - let (_, (_, t)) = Tac2env.interp_global kn in - Feedback.msg_notice (pr_qualid qid ++ spc () ++ str ":" ++ spc () ++ pr_glbtype t) + match kn with + | TacConstant kn -> + let (_, (_, t)) = Tac2env.interp_global kn in + Feedback.msg_notice (pr_qualid qid ++ spc () ++ str ":" ++ spc () ++ pr_glbtype t) + | TacConstructor kn -> + let _ = Tac2env.interp_constructor kn in + Feedback.msg_notice (str "Constructor" ++ spc () ++ str ":" ++ spc () ++ pr_qualid qid) (** Calling tactics *) diff --git a/tac2env.ml b/tac2env.ml index 3500b2ef3d..bdb8f41ef8 100644 --- a/tac2env.ml +++ b/tac2env.ml @@ -15,23 +15,35 @@ open Tac2expr type ltac_constant = KerName.t type ltac_constructor = KerName.t +type ltac_projection = KerName.t type type_constant = KerName.t type constructor_data = { + cdata_prms : int; cdata_type : type_constant; - cdata_args : int * int glb_typexpr list; + cdata_args : int glb_typexpr list; cdata_indx : int; } +type projection_data = { + pdata_prms : int; + pdata_type : type_constant; + pdata_ptyp : int glb_typexpr; + pdata_mutb : bool; + pdata_indx : int; +} + type ltac_state = { ltac_tactics : (glb_tacexpr * type_scheme) KNmap.t; ltac_constructors : constructor_data KNmap.t; + ltac_projections : projection_data KNmap.t; ltac_types : glb_quant_typedef KNmap.t; } let empty_state = { ltac_tactics = KNmap.empty; ltac_constructors = KNmap.empty; + ltac_projections = KNmap.empty; ltac_types = KNmap.empty; } @@ -51,8 +63,8 @@ let rec eval_pure = function | GTacTup el -> ValBlk (0, Array.map_of_list eval_pure el) | GTacCst (_, n, []) -> ValInt n | GTacCst (_, n, el) -> ValBlk (n, Array.map_of_list eval_pure el) -| GTacAtm (AtmStr _) | GTacArr _ | GTacLet _ | GTacVar _ -| GTacApp _ | GTacCse _ | GTacPrm _ | GTacExt _ -> +| GTacAtm (AtmStr _) | GTacArr _ | GTacLet _ | GTacVar _ | GTacSet _ +| GTacApp _ | GTacCse _ | GTacPrj _ | GTacPrm _ | GTacExt _ -> anomaly (Pp.str "Term is not a syntactical value") let define_global kn e = @@ -69,6 +81,12 @@ let define_constructor kn t = let interp_constructor kn = KNmap.find kn ltac_state.contents.ltac_constructors +let define_projection kn t = + let state = !ltac_state in + ltac_state := { state with ltac_projections = KNmap.add kn t state.ltac_projections } + +let interp_projection kn = KNmap.find kn ltac_state.contents.ltac_projections + let define_type kn e = let state = !ltac_state in ltac_state := { state with ltac_types = KNmap.add kn e state.ltac_types } @@ -103,28 +121,47 @@ struct id, (DirPath.repr dir) end +type tacref = +| TacConstant of ltac_constant +| TacConstructor of ltac_constructor + +module TacRef = +struct +type t = tacref +let equal r1 r2 = match r1, r2 with +| TacConstant c1, TacConstant c2 -> KerName.equal c1 c2 +| TacConstructor c1, TacConstructor c2 -> KerName.equal c1 c2 +| _ -> false +end + module KnTab = Nametab.Make(FullPath)(KerName) +module RfTab = Nametab.Make(FullPath)(TacRef) type nametab = { - tab_ltac : KnTab.t; + tab_ltac : RfTab.t; tab_type : KnTab.t; + tab_proj : KnTab.t; } -let empty_nametab = { tab_ltac = KnTab.empty; tab_type = KnTab.empty } +let empty_nametab = { + tab_ltac = RfTab.empty; + tab_type = KnTab.empty; + tab_proj = KnTab.empty; +} let nametab = Summary.ref empty_nametab ~name:"ltac2-nametab" let push_ltac vis sp kn = let tab = !nametab in - nametab := { tab with tab_ltac = KnTab.push vis sp kn tab.tab_ltac } + nametab := { tab with tab_ltac = RfTab.push vis sp kn tab.tab_ltac } let locate_ltac qid = let tab = !nametab in - KnTab.locate qid tab.tab_ltac + RfTab.locate qid tab.tab_ltac let locate_extended_all_ltac qid = let tab = !nametab in - KnTab.find_prefixes qid tab.tab_ltac + RfTab.find_prefixes qid tab.tab_ltac let push_type vis sp kn = let tab = !nametab in @@ -138,6 +175,18 @@ let locate_extended_all_type qid = let tab = !nametab in KnTab.find_prefixes qid tab.tab_type +let push_projection vis sp kn = + let tab = !nametab in + nametab := { tab with tab_proj = KnTab.push vis sp kn tab.tab_proj } + +let locate_projection qid = + let tab = !nametab in + KnTab.locate qid tab.tab_proj + +let locate_extended_all_projection qid = + let tab = !nametab in + KnTab.find_prefixes qid tab.tab_proj + type 'a ml_object = { ml_type : type_constant; ml_interp : environment -> 'a -> Geninterp.Val.t Proofview.tactic; diff --git a/tac2env.mli b/tac2env.mli index eb471b9abf..bcfa70487a 100644 --- a/tac2env.mli +++ b/tac2env.mli @@ -14,10 +14,6 @@ open Tac2expr (** Ltac2 global environment *) -type ltac_constant = KerName.t -type ltac_constructor = KerName.t -type type_constant = KerName.t - (** {5 Toplevel definition of values} *) val define_global : ltac_constant -> (glb_tacexpr * type_scheme) -> unit @@ -31,24 +27,57 @@ val interp_type : type_constant -> glb_quant_typedef (** {5 Toplevel definition of algebraic constructors} *) type constructor_data = { + cdata_prms : int; + (** Type parameters *) cdata_type : type_constant; - cdata_args : int * int glb_typexpr list; + (** Inductive definition to which the constructor pertains *) + cdata_args : int glb_typexpr list; + (** Types of the constructor arguments *) cdata_indx : int; + (** Index of the constructor in the ADT. Numbering is duplicated between + argumentless and argument-using constructors, e.g. in type ['a option] + [None] and [Some] have both index 0. *) } val define_constructor : ltac_constructor -> constructor_data -> unit val interp_constructor : ltac_constructor -> constructor_data +(** {5 Toplevel definition of projections} *) + +type projection_data = { + pdata_prms : int; + (** Type parameters *) + pdata_type : type_constant; + (** Record definition to which the projection pertains *) + pdata_ptyp : int glb_typexpr; + (** Type of the projection *) + pdata_mutb : bool; + (** Whether the field is mutable *) + pdata_indx : int; + (** Index of the projection *) +} + +val define_projection : ltac_projection -> projection_data -> unit +val interp_projection : ltac_projection -> projection_data + (** {5 Name management} *) -val push_ltac : visibility -> full_path -> ltac_constant -> unit -val locate_ltac : qualid -> ltac_constant -val locate_extended_all_ltac : qualid -> ltac_constant list +type tacref = +| TacConstant of ltac_constant +| TacConstructor of ltac_constructor + +val push_ltac : visibility -> full_path -> tacref -> unit +val locate_ltac : qualid -> tacref +val locate_extended_all_ltac : qualid -> tacref list val push_type : visibility -> full_path -> type_constant -> unit val locate_type : qualid -> type_constant val locate_extended_all_type : qualid -> type_constant list +val push_projection : visibility -> full_path -> ltac_projection -> unit +val locate_projection : qualid -> ltac_projection +val locate_extended_all_projection : qualid -> ltac_projection list + (** {5 Toplevel definitions of ML tactics} *) (** This state is not part of the summary, contrarily to the ones above. It is diff --git a/tac2expr.mli b/tac2expr.mli index 15c630ca87..b9b649e481 100644 --- a/tac2expr.mli +++ b/tac2expr.mli @@ -16,6 +16,11 @@ type rec_flag = bool type lid = Id.t type uid = Id.t +type ltac_constant = KerName.t +type ltac_constructor = KerName.t +type ltac_projection = KerName.t +type type_constant = KerName.t + (** {5 Misc} *) type ml_tactic_name = { @@ -40,7 +45,7 @@ type 'a glb_typexpr = | GTypVar of 'a | GTypArrow of 'a glb_typexpr * 'a glb_typexpr | GTypTuple of 'a glb_typexpr list -| GTypRef of KerName.t * 'a glb_typexpr list +| GTypRef of type_constant * 'a glb_typexpr list type glb_typedef = | GTydDef of int glb_typexpr option @@ -76,25 +81,32 @@ type raw_tacexpr = | CTacCnv of Loc.t * raw_tacexpr * raw_typexpr | CTacSeq of Loc.t * raw_tacexpr * raw_tacexpr | CTacCse of Loc.t * raw_tacexpr * raw_taccase list +| CTacRec of Loc.t * raw_recexpr +| CTacPrj of Loc.t * raw_tacexpr * qualid located +| CTacSet of Loc.t * raw_tacexpr * qualid located * raw_tacexpr | CTacExt of Loc.t * raw_generic_argument and raw_taccase = raw_patexpr * raw_tacexpr +and raw_recexpr = (qualid located * raw_tacexpr) list + type case_info = | GCaseTuple of int -| GCaseAlg of KerName.t +| GCaseAlg of type_constant type glb_tacexpr = | GTacAtm of atom | GTacVar of Id.t -| GTacRef of KerName.t +| GTacRef of ltac_constant | GTacFun of Name.t list * glb_tacexpr | GTacApp of glb_tacexpr * glb_tacexpr list | GTacLet of rec_flag * (Name.t * glb_tacexpr) list * glb_tacexpr | GTacTup of glb_tacexpr list | GTacArr of glb_tacexpr list -| GTacCst of KerName.t * int * glb_tacexpr list +| GTacCst of type_constant * int * glb_tacexpr list | GTacCse of glb_tacexpr * case_info * glb_tacexpr array * (Name.t array * glb_tacexpr) array +| GTacPrj of glb_tacexpr * int +| GTacSet of glb_tacexpr * int * glb_tacexpr | GTacExt of glob_generic_argument | GTacPrm of ml_tactic_name * glb_tacexpr list diff --git a/tac2intern.ml b/tac2intern.ml index a554f959a0..10fcde6efa 100644 --- a/tac2intern.ml +++ b/tac2intern.ml @@ -174,6 +174,9 @@ let loc_of_tacexpr = function | CTacCnv (loc, _, _) -> loc | CTacSeq (loc, _, _) -> loc | CTacCse (loc, _, _) -> loc +| CTacRec (loc, _) -> loc +| CTacPrj (loc, _, _) -> loc +| CTacSet (loc, _, _, _) -> loc | CTacExt (loc, _) -> loc let loc_of_patexpr = function @@ -181,6 +184,11 @@ let loc_of_patexpr = function | CPatRef (loc, _, _) -> loc | CPatTup (loc, _) -> loc +let error_nargs_mismatch loc nargs nfound = + user_err ~loc (str "Constructor expects " ++ int nargs ++ + str " arguments, but is applied to " ++ int nfound ++ + str " arguments") + let rec subst_type subst (t : 'a glb_typexpr) = match t with | GTypVar id -> subst id | GTypArrow (t1, t2) -> GTypArrow (subst_type subst t1, subst_type subst t2) @@ -320,20 +328,27 @@ let unify loc env t1 t2 = (** Term typing *) +let is_pure_constructor kn = + match snd (Tac2env.interp_type kn) with + | GTydAlg _ -> true + | GTydRec fields -> + let is_pure (_, mut, _) = not mut in + List.for_all is_pure fields + | GTydDef _ -> assert false (** Type definitions have no constructors *) + let rec is_value = function | GTacAtm (AtmInt _) | GTacVar _ | GTacRef _ | GTacFun _ -> true | GTacAtm (AtmStr _) | GTacApp _ | GTacLet _ -> false | GTacTup el -> List.for_all is_value el | GTacCst (_, _, []) -> true -| GTacCst (kn, n, el) -> - (** To be a value, a constructor must be immutable *) - assert false (** TODO *) -| GTacArr _ | GTacCse _ | GTacExt _ | GTacPrm _ -> false +| GTacCst (kn, _, el) -> is_pure_constructor kn && List.for_all is_value el +| GTacArr _ | GTacCse _ | GTacPrj _ | GTacSet _ | GTacExt _ | GTacPrm _ -> false let is_rec_rhs = function | GTacFun _ -> true -| GTacAtm _ | GTacVar _ | GTacRef _ | GTacApp _ | GTacLet _ -> false -| GTacTup _ | GTacArr _ | GTacExt _ | GTacPrm _ | GTacCst _ | GTacCse _ -> false +| GTacAtm _ | GTacVar _ | GTacRef _ | GTacApp _ | GTacLet _ | GTacPrj _ +| GTacSet _ | GTacTup _ | GTacArr _ | GTacExt _ | GTacPrm _ | GTacCst _ +| GTacCse _ -> false let rec fv_type f t accu = match t with | GTypVar id -> f id accu @@ -438,18 +453,23 @@ let get_variable env (loc, qid) = let get_constructor env (loc, qid) = let c = try Some (Tac2env.locate_ltac qid) with Not_found -> None in match c with - | Some knc -> - let kn = - try Tac2env.interp_constructor knc - with Not_found -> - CErrors.user_err ~loc (str "The term " ++ pr_qualid qid ++ - str " is not the constructor of an inductive type.") in + | Some (TacConstructor knc) -> + let kn = Tac2env.interp_constructor knc in ArgArg (kn, knc) + | Some (TacConstant _) -> + CErrors.user_err ~loc (str "The term " ++ pr_qualid qid ++ + str " is not the constructor of an inductive type.") | None -> let (dp, id) = repr_qualid qid in if DirPath.is_empty dp then ArgVar (loc, id) else CErrors.user_err ~loc (str "Unbound constructor " ++ pr_qualid qid) +let get_projection (loc, qid) = + let kn = try Tac2env.locate_projection qid with Not_found -> + user_err ~loc (pr_qualid qid ++ str " is not a projection") + in + Tac2env.interp_projection kn + let intern_atm env = function | AtmInt n -> (GTacAtm (AtmInt n), GTypRef (t_int, [])) | AtmStr s -> (GTacAtm (AtmStr s), GTypRef (t_string, [])) @@ -496,6 +516,10 @@ let get_pattern_kind env pl = match pl with in get_kind p pl +let is_constructor env qid = match get_variable env qid with +| ArgArg (TacConstructor _) -> true +| _ -> false + let rec intern_rec env = function | CTacAtm (_, atm) -> intern_atm env atm | CTacRef qid -> @@ -503,9 +527,11 @@ let rec intern_rec env = function | ArgVar (_, id) -> let sch = Id.Map.find id env.env_var in (GTacVar id, fresh_mix_type_scheme env sch) - | ArgArg kn -> + | ArgArg (TacConstant kn) -> let (_, sch) = Tac2env.interp_global kn in (GTacRef kn, fresh_type_scheme env sch) + | ArgArg (TacConstructor kn) -> + intern_constructor env (fst qid) kn [] end | CTacFun (loc, bnd, e) -> let fold (env, bnd, tl) ((_, na), t) = @@ -520,6 +546,12 @@ let rec intern_rec env = function let (e, t) = intern_rec env e in let t = List.fold_left (fun accu t -> GTypArrow (t, accu)) t tl in (GTacFun (bnd, e), t) +| CTacApp (loc, CTacRef qid, args) when is_constructor env qid -> + let kn = match get_variable env qid with + | ArgArg (TacConstructor kn) -> kn + | _ -> assert false + in + intern_constructor env (fst qid) kn args | CTacApp (loc, f, args) -> let (f, ft) = intern_rec env f in let fold arg (args, t) = @@ -571,12 +603,7 @@ let rec intern_rec env = function (GTacArr [], GTypRef (t_int, [GTypVar id])) | CTacArr (loc, e0 :: el) -> let (e0, t0) = intern_rec env e0 in - let fold e el = - let loc = loc_of_tacexpr e in - let (e, t) = intern_rec env e in - let () = unify loc env t t0 in - e :: el - in + let fold e el = intern_rec_with_constraint env e t0 :: el in let el = e0 :: List.fold_right fold el [] in (GTacArr el, GTypRef (t_array, [t0])) | CTacLst (loc, []) -> @@ -584,12 +611,7 @@ let rec intern_rec env = function (c_nil, GTypRef (t_list, [GTypVar id])) | CTacLst (loc, e0 :: el) -> let (e0, t0) = intern_rec env e0 in - let fold e el = - let loc = loc_of_tacexpr e in - let (e, t) = intern_rec env e in - let () = unify loc env t t0 in - c_cons e el - in + let fold e el = c_cons (intern_rec_with_constraint env e t0) el in let el = c_cons e0 (List.fold_right fold el c_nil) in (el, GTypRef (t_list, [t0])) | CTacCnv (loc, e, tc) -> @@ -604,6 +626,34 @@ let rec intern_rec env = function (GTacLet (false, [Anonymous, e1], e2), t2) | CTacCse (loc, e, pl) -> intern_case env loc e pl +| CTacRec (loc, fs) -> + intern_record env loc fs +| CTacPrj (loc, e, proj) -> + let pinfo = get_projection proj in + let loc = loc_of_tacexpr e in + let (e, t) = intern_rec env e in + let subst = Array.init pinfo.pdata_prms (fun _ -> fresh_id env) in + let params = Array.map_to_list (fun i -> GTypVar i) subst in + let exp = GTypRef (pinfo.pdata_type, params) in + let () = unify loc env t exp in + let substf i = GTypVar subst.(i) in + let ret = subst_type substf pinfo.pdata_ptyp in + (GTacPrj (e, pinfo.pdata_indx), ret) +| CTacSet (loc, e, proj, r) -> + let pinfo = get_projection proj in + let () = + if not pinfo.pdata_mutb then + let (loc, _) = proj in + user_err ~loc (str "Field is not mutable") + in + let subst = Array.init pinfo.pdata_prms (fun _ -> fresh_id env) in + let params = Array.map_to_list (fun i -> GTypVar i) subst in + let exp = GTypRef (pinfo.pdata_type, params) in + let e = intern_rec_with_constraint env e exp in + let substf i = GTypVar subst.(i) in + let ret = subst_type substf pinfo.pdata_ptyp in + let r = intern_rec_with_constraint env r ret in + (GTacSet (e, pinfo.pdata_indx, r), GTypRef (t_unit, [])) | CTacExt (loc, ext) -> let open Genintern in let GenArg (Rawwit tag, _) = ext in @@ -616,6 +666,12 @@ let rec intern_rec env = function let (_, ext) = Flags.with_option Ltac_plugin.Tacintern.strict_check (fun () -> generic_intern ist ext) () in (GTacExt ext, GTypRef (tpe.ml_type, [])) +and intern_rec_with_constraint env e exp = + let loc = loc_of_tacexpr e in + let (e, t) = intern_rec env e in + let () = unify loc env t exp in + e + and intern_let_rec env loc el e = let fold accu ((loc, na), _, _) = match na with | Anonymous -> accu @@ -760,12 +816,9 @@ and intern_case env loc e pl = in let ids = List.map get_id args in let nids = List.length ids in - let nargs = List.length (snd data.cdata_args) in + let nargs = List.length data.cdata_args in let () = - if not (Int.equal nids nargs) then - user_err ~loc (str "Constructor expects " ++ int nargs ++ - str " arguments, but is applied to " ++ int nids ++ - str " arguments") + if not (Int.equal nids nargs) then error_nargs_mismatch loc nargs nids in let fold env id tpe = (** Instantiate all arguments *) @@ -773,7 +826,7 @@ and intern_case env loc e pl = let tpe = subst_type subst tpe in push_name id (monomorphic tpe) env in - let nenv = List.fold_left2 fold env ids (snd data.cdata_args) in + let nenv = List.fold_left2 fold env ids data.cdata_args in let (br', brT) = intern_rec nenv br in let () = let index = data.cdata_indx in @@ -802,6 +855,64 @@ and intern_case env loc e pl = let ce = GTacCse (e', GCaseAlg kn, const, nonconst) in (ce, ret) +and intern_constructor env loc kn args = + let cstr = interp_constructor kn in + let nargs = List.length cstr.cdata_args in + if Int.equal nargs (List.length args) then + let subst = Array.init cstr.cdata_prms (fun _ -> fresh_id env) in + let substf i = GTypVar subst.(i) in + let types = List.map (fun t -> subst_type substf t) cstr.cdata_args in + let ans = GTypRef (cstr.cdata_type, List.init cstr.cdata_prms (fun i -> GTypVar subst.(i))) in + let map arg tpe = intern_rec_with_constraint env arg tpe in + let args = List.map2 map args types in + (GTacCst (cstr.cdata_type, cstr.cdata_indx, args), ans) + else + error_nargs_mismatch loc nargs (List.length args) + +and intern_record env loc fs = + let map ((loc, qid), e) = + let proj = get_projection (loc, qid) in + (loc, proj, e) + in + let fs = List.map map fs in + let kn = match fs with + | [] -> user_err ~loc (str "Cannot infer the corresponding record type") + | (_, proj, _) :: _ -> proj.pdata_type + in + let params, typdef = match Tac2env.interp_type kn with + | n, GTydRec def -> n, def + | _ -> assert false + in + let subst = Array.init params (fun _ -> fresh_id env) in + (** Set the answer [args] imperatively *) + let args = Array.make (List.length typdef) None in + let iter (loc, pinfo, e) = + if KerName.equal kn pinfo.pdata_type then + let index = pinfo.pdata_indx in + match args.(index) with + | None -> + let exp = subst_type (fun i -> GTypVar subst.(i)) pinfo.pdata_ptyp in + let e = intern_rec_with_constraint env e exp in + args.(index) <- Some e + | Some _ -> + let (name, _, _) = List.nth typdef pinfo.pdata_indx in + user_err ~loc (str "Field " ++ Id.print name ++ str " is defined \ + several times") + else + user_err ~loc (str "Field " ++ (*KerName.print knp ++*) str " does not \ + pertain to record definition " ++ KerName.print pinfo.pdata_type) + in + let () = List.iter iter fs in + let () = match Array.findi (fun _ o -> Option.is_empty o) args with + | None -> () + | Some i -> + let (field, _, _) = List.nth typdef i in + user_err ~loc (str "Field " ++ Id.print field ++ str " is undefined") + in + let args = Array.map_to_list Option.get args in + let tparam = List.init params (fun i -> GTypVar subst.(i)) in + (GTacCst (kn, 0, args), GTypRef (kn, tparam)) + let normalize env (count, vars) (t : UF.elt glb_typexpr) = let get_var id = try UF.Map.find id !vars @@ -907,6 +1018,13 @@ let rec subst_expr subst e = match e with let cse1' = Array.map (fun (ids, e) -> (ids, subst_expr subst e)) cse1 in let ci' = subst_case_info subst ci in GTacCse (subst_expr subst e, ci', cse0', cse1') +| GTacPrj (e, p) as e0 -> + let e' = subst_expr subst e in + if e' == e then e0 else GTacPrj (e', p) +| GTacSet (e, p, r) as e0 -> + let e' = subst_expr subst e in + let r' = subst_expr subst r in + if e' == e && r' == r then e0 else GTacSet (e', p, r') | GTacExt ext -> let ext' = Genintern.generic_substitute subst ext in if ext' == ext then e else GTacExt ext' diff --git a/tac2interp.ml b/tac2interp.ml index b868caf963..fedbb13e7d 100644 --- a/tac2interp.ml +++ b/tac2interp.ml @@ -80,6 +80,12 @@ let rec interp ist = function return (ValBlk (n, Array.of_list el)) | GTacCse (e, _, cse0, cse1) -> interp ist e >>= fun e -> interp_case ist e cse0 cse1 +| GTacPrj (e, p) -> + interp ist e >>= fun e -> interp_proj ist e p +| GTacSet (e, p, r) -> + interp ist e >>= fun e -> + interp ist r >>= fun r -> + interp_set ist e p r | GTacPrm (ml, el) -> Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> Tac2env.interp_primitive ml el @@ -110,3 +116,16 @@ and interp_case ist e cse0 cse1 = match e with interp ist e | ValExt _ | ValStr _ | ValCls _ -> anomaly (str "Unexpected value shape") + +and interp_proj ist e p = match e with +| ValBlk (_, args) -> + return args.(p) +| ValInt _ | ValExt _ | ValStr _ | ValCls _ -> + anomaly (str "Unexpected value shape") + +and interp_set ist e p r = match e with +| ValBlk (_, args) -> + let () = args.(p) <- r in + return (ValInt 0) +| ValInt _ | ValExt _ | ValStr _ | ValCls _ -> + anomaly (str "Unexpected value shape") -- cgit v1.2.3 From 152b259b7b587ea949dd856b24beaf56466f3f27 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 5 Dec 2016 16:32:26 +0100 Subject: Fixing a precedence issue in type parameters. --- Array.v | 14 +++ Constr.v | 9 ++ Control.v | 44 +++++++ Init.v | 32 +++++ Int.v | 16 +++ Ltac2.v | 42 ++----- Message.v | 20 +++ String.v | 14 +++ g_ltac2.ml4 | 14 ++- tac2core.ml | 381 ++++++++++++++++++++++++++++++++++++++++++++++++++------- tac2core.mli | 19 +++ tac2entries.ml | 3 +- tac2env.ml | 5 + tac2env.mli | 5 + tac2expr.mli | 2 +- tac2intern.ml | 3 +- tac2interp.ml | 10 +- tac2interp.mli | 12 +- vo.itarget | 7 ++ 19 files changed, 556 insertions(+), 96 deletions(-) create mode 100644 Array.v create mode 100644 Constr.v create mode 100644 Control.v create mode 100644 Init.v create mode 100644 Int.v create mode 100644 Message.v create mode 100644 String.v diff --git a/Array.v b/Array.v new file mode 100644 index 0000000000..be4ab025ae --- /dev/null +++ b/Array.v @@ -0,0 +1,14 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* 'a -> 'a array := "ltac2" "array_make". +Ltac2 @external length : 'a array -> int := "ltac2" "array_length". +Ltac2 @external get : 'a array -> int -> 'a := "ltac2" "array_get". +Ltac2 @external set : 'a array -> int -> 'a -> unit := "ltac2" "array_set". diff --git a/Constr.v b/Constr.v new file mode 100644 index 0000000000..994f9bf3ac --- /dev/null +++ b/Constr.v @@ -0,0 +1,9 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* 'a := "ltac2" "throw". +(** Fatal exception throwing. This does not induce backtracking. *) + +(** Generic backtracking control *) + +Ltac2 @ external zero : exn -> 'a := "ltac2" "zero". +Ltac2 @ external plus : (unit -> 'a) -> (exn -> 'a) -> 'a := "ltac2" "plus". +Ltac2 @ external once : (unit -> 'a) -> 'a := "ltac2" "once". +Ltac2 @ external dispatch : (unit -> unit) list -> unit := "ltac2" "dispatch". +Ltac2 @ external extend : (unit -> unit) list -> (unit -> unit) -> (unit -> unit) list -> unit := "ltac2" "extend". +Ltac2 @ external enter : (unit -> unit) -> unit := "ltac2" "enter". + +(** Proof state manipulation *) + +Ltac2 @ external focus : int -> int -> (unit -> 'a) -> 'a := "ltac2" "focus". +Ltac2 @ external shelve : unit -> unit := "ltac2" "shelve". +Ltac2 @ external shelve_unifiable : unit -> unit := "ltac2" "shelve_unifiable". + +(** Goal inspection *) + +Ltac2 @ external goal : unit -> constr := "ltac2" "goal". +(** Panics if there is not exactly one goal under focus. Otherwise returns + the conclusion of this goal. *) + +Ltac2 @ external hyp : ident -> constr := "ltac2" "hyp". +(** Panics if there is more than one goal under focus. If there is no + goal under focus, looks for the section variable with the given name. + If there is one, looks for the hypothesis with the given name. *) + +(** Refinement *) + +Ltac2 @ external refine : (unit -> constr) -> unit := "ltac2" "refine". diff --git a/Init.v b/Init.v new file mode 100644 index 0000000000..0f0b4636d8 --- /dev/null +++ b/Init.v @@ -0,0 +1,32 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* int -> bool := "ltac2" "int_equal". +Ltac2 @ external compare : int -> int -> int := "ltac2" "int_compare". +Ltac2 @ external add : int -> int -> int := "ltac2" "int_add". +Ltac2 @ external sub : int -> int -> int := "ltac2" "int_sub". +Ltac2 @ external mul : int -> int -> int := "ltac2" "int_mul". +Ltac2 @ external neg : int -> int := "ltac2" "int_neg". diff --git a/Ltac2.v b/Ltac2.v index 0933c1e0b4..625d4ac513 100644 --- a/Ltac2.v +++ b/Ltac2.v @@ -6,37 +6,11 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -Declare ML Module "ltac2_plugin". - -Global Set Default Proof Mode "Ltac2". - -(** Primitive types *) - -Ltac2 Type int. -Ltac2 Type string. -Ltac2 Type constr. -Ltac2 Type message. -Ltac2 Type 'a array. - -(** Pervasive types *) - -Ltac2 Type 'a option := [ None | Some ('a) ]. - -(** Primitive tactics *) - -Module Message. - -Ltac2 @ external print : message -> unit := "ltac2" "print". -Ltac2 @ external of_string : string -> message := "ltac2" "message_of_string". -Ltac2 @ external of_int : int -> message := "ltac2" "message_of_int". - -End Message. - -Module Array. - -Ltac2 @external make : int -> 'a -> ('a) array := "ltac2" "array_make". -Ltac2 @external length : ('a) array -> int := "ltac2" "array_length". -Ltac2 @external get : ('a) array -> int -> 'a := "ltac2" "array_get". -Ltac2 @external set : ('a) array -> int -> 'a -> unit := "ltac2" "array_set". - -End Array. +Require Export Coq.ltac2.Init. + +Require Coq.ltac2.Int. +Require Coq.ltac2.String. +Require Coq.ltac2.Array. +Require Coq.ltac2.Message. +Require Coq.ltac2.Constr. +Require Coq.ltac2.Control. diff --git a/Message.v b/Message.v new file mode 100644 index 0000000000..36233f4544 --- /dev/null +++ b/Message.v @@ -0,0 +1,20 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* unit := "ltac2" "print". + +Ltac2 @ external of_string : string -> message := "ltac2" "message_of_string". + +Ltac2 @ external of_int : int -> message := "ltac2" "message_of_int". + +Ltac2 @ external of_constr : constr -> message := "ltac2" "message_of_constr". +(** Panics if there is more than one goal under focus. *) + +Ltac2 @ external concat : message -> message -> message := "ltac2" "message_concat". diff --git a/String.v b/String.v new file mode 100644 index 0000000000..3a4a178878 --- /dev/null +++ b/String.v @@ -0,0 +1,14 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* char -> string := "ltac2" "string_make". +Ltac2 @external length : string -> int := "ltac2" "string_length". +Ltac2 @external get : string -> int -> char := "ltac2" "string_get". +Ltac2 @external set : string -> int -> char -> unit := "ltac2" "string_set". diff --git a/g_ltac2.ml4 b/g_ltac2.ml4 index ce2becd9f9..ff3d79bbae 100644 --- a/g_ltac2.ml4 +++ b/g_ltac2.ml4 @@ -23,6 +23,8 @@ let tac2mode = Gram.entry_create "vernac:ltac2_command" let inj_wit wit loc x = CTacExt (loc, Genarg.in_gen (Genarg.rawwit wit) x) let inj_constr loc c = inj_wit Stdarg.wit_constr loc c +let inj_open_constr loc c = inj_wit Stdarg.wit_open_constr loc c +let inj_ident loc c = inj_wit Stdarg.wit_ident loc c GEXTEND Gram GLOBAL: tac2expr tac2type tac2def_val tac2def_typ tac2def_ext; @@ -84,6 +86,8 @@ GEXTEND Gram | s = Prim.string -> CTacAtm (!@loc, AtmStr s) | id = Prim.qualid -> CTacRef id | IDENT "constr"; ":"; "("; c = Constr.lconstr; ")" -> inj_constr !@loc c + | IDENT "open_constr"; ":"; "("; c = Constr.lconstr; ")" -> inj_open_constr !@loc c + | IDENT "ident"; ":"; "("; c = Prim.ident; ")" -> inj_ident !@loc c ] ] ; let_clause: @@ -97,13 +101,15 @@ GEXTEND Gram [ t1 = tac2type; "->"; t2 = tac2type -> CTypArrow (!@loc, t1, t2) ] | "2" [ t = tac2type; "*"; tl = LIST1 tac2type SEP "*" -> CTypTuple (!@loc, t :: tl) ] - | "1" - [ "("; p = LIST1 tac2type LEVEL "5" SEP ","; ")"; qid = Prim.qualid -> CTypRef (!@loc, qid, p) ] + | "1" LEFTA + [ t = SELF; qid = Prim.qualid -> CTypRef (!@loc, qid, [t]) ] | "0" - [ "("; t = tac2type; ")" -> t + [ "("; t = tac2type LEVEL "5"; ")" -> t | id = typ_param -> CTypVar (!@loc, Name id) | "_" -> CTypVar (!@loc, Anonymous) - | qid = Prim.qualid -> CTypRef (!@loc, qid, []) ] + | qid = Prim.qualid -> CTypRef (!@loc, qid, []) + | "("; p = LIST1 tac2type LEVEL "5" SEP ","; ")"; qid = Prim.qualid -> + CTypRef (!@loc, qid, p) ] ]; locident: [ [ id = Prim.ident -> (!@loc, id) ] ] diff --git a/tac2core.ml b/tac2core.ml index 95632bf7b1..fd998151fd 100644 --- a/tac2core.ml +++ b/tac2core.ml @@ -14,12 +14,25 @@ open Genarg open Geninterp open Tac2env open Tac2expr +open Tac2interp open Proofview.Notations (** Standard values *) -let coq_prefix = DirPath.make (List.map Id.of_string ["Ltac2"; "ltac2"; "Coq"]) -let coq_core n = KerName.make2 (MPfile coq_prefix) (Label.of_id (Id.of_string_soft n)) +let coq_core n = KerName.make2 Tac2env.coq_prefix (Label.of_id (Id.of_string_soft n)) + +let val_tag t = match val_tag t with +| Val.Base t -> t +| _ -> assert false + +let val_constr = val_tag (topwit Stdarg.wit_constr) +let val_ident = val_tag (topwit Stdarg.wit_ident) +let val_pp = Val.create "ltac2:pp" + +let extract_val (type a) (tag : a Val.typ) (Val.Dyn (tag', v)) : a = +match Val.eq tag tag' with +| None -> assert false +| Some Refl -> v module Core = struct @@ -30,6 +43,7 @@ let t_array = coq_core "array" let t_unit = coq_core "unit" let t_list = coq_core "list" let t_constr = coq_core "constr" +let t_ident = coq_core "ident" let c_nil = coq_core "[]" let c_cons = coq_core "::" @@ -51,6 +65,28 @@ let to_unit = function | ValInt 0 -> () | _ -> assert false +let of_int n = ValInt n +let to_int = function +| ValInt n -> n +| _ -> assert false + +let of_bool b = if b then ValInt 0 else ValInt 1 + +let to_bool = function +| ValInt 0 -> true +| ValInt 1 -> false +| _ -> assert false + +let of_char n = ValInt (Char.code n) +let to_char = function +| ValInt n -> Char.chr n +| _ -> assert false + +let of_string s = ValStr s +let to_string = function +| ValStr s -> s +| _ -> assert false + let rec of_list = function | [] -> v_nil | x :: l -> v_cons x (of_list l) @@ -60,20 +96,33 @@ let rec to_list = function | ValBlk (0, [|v; vl|]) -> v :: to_list vl | _ -> assert false -end +let of_ext tag c = + ValExt (Val.Dyn (tag, c)) -let extract_val (type a) (tag : a Val.typ) (Val.Dyn (tag', v)) : a = -match Val.eq tag tag' with -| None -> assert false -| Some Refl -> v +let to_ext tag = function +| ValExt e -> extract_val tag e +| _ -> assert false -let val_pp = Val.create "ltac2:pp" -let get_pp v = extract_val val_pp v +let of_constr c = of_ext val_constr c +let to_constr c = to_ext val_constr c + +let of_ident c = of_ext val_ident c +let to_ident c = to_ext val_ident c + +let of_exn c = of_ext val_exn c +let to_exn c = to_ext val_exn c + +let of_pp c = of_ext val_pp c +let to_pp c = to_ext val_pp c + +end let val_valexpr = Val.create "ltac2:valexpr" (** Helper functions *) +let thaw f = interp_app f [v_unit] + let return x = Proofview.tclUNIT x let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } @@ -83,10 +132,35 @@ let wrap f = let wrap_unit f = return () >>= fun () -> f (); return v_unit +(** In Ltac2, the notion of "current environment" only makes sense when there is + at most one goal under focus. Contrarily to Ltac1, instead of dynamically + focussing when we need it, we raise a non-backtracking error when it does + not make sense. *) +exception NonFocussedGoal + +let () = register_handler begin function +| NonFocussedGoal -> str "Several goals under focus" +| _ -> raise Unhandled +end + +let pf_apply f = + Proofview.Goal.goals >>= function + | [] -> + Proofview.tclENV >>= fun env -> + Proofview.tclEVARMAP >>= fun sigma -> + f env sigma + | [gl] -> + gl >>= fun gl -> + f (Proofview.Goal.env gl) (Tacmach.New.project gl) + | _ :: _ :: _ -> + Proofview.tclLIFT (Proofview.NonLogical.raise NonFocussedGoal) + (** Primitives *) +(** Printing *) + let prm_print : ml_tactic = function -| [ValExt pp] -> wrap_unit (fun () -> Feedback.msg_notice (get_pp pp)) +| [pp] -> wrap_unit (fun () -> Feedback.msg_notice (Value.to_pp pp)) | _ -> assert false let prm_message_of_int : ml_tactic = function @@ -94,11 +168,31 @@ let prm_message_of_int : ml_tactic = function | _ -> assert false let prm_message_of_string : ml_tactic = function -| [ValStr s] -> return (ValExt (Val.Dyn (val_pp, str s))) +| [ValStr s] -> return (ValExt (Val.Dyn (val_pp, str (Bytes.to_string s)))) +| _ -> assert false + +let prm_message_of_constr : ml_tactic = function +| [c] -> + pf_apply begin fun env sigma -> + let c = Value.to_constr c in + let pp = Printer.pr_econstr_env env sigma c in + return (ValExt (Val.Dyn (val_pp, pp))) + end | _ -> assert false +let prm_message_concat : ml_tactic = function +| [m1; m2] -> + let m1 = Value.to_pp m1 in + let m2 = Value.to_pp m2 in + return (Value.of_pp (Pp.app m1 m2)) +| _ -> assert false + +(** Array *) + let prm_array_make : ml_tactic = function -| [ValInt n; x] -> wrap (fun () -> ValBlk (0, Array.make n x)) +| [ValInt n; x] -> + (** FIXME: wrap exception *) + wrap (fun () -> ValBlk (0, Array.make n x)) | _ -> assert false let prm_array_length : ml_tactic = function @@ -106,31 +200,214 @@ let prm_array_length : ml_tactic = function | _ -> assert false let prm_array_set : ml_tactic = function -| [ValBlk (_, v); ValInt n; x] -> wrap_unit (fun () -> v.(n) <- x) +| [ValBlk (_, v); ValInt n; x] -> + (** FIXME: wrap exception *) + wrap_unit (fun () -> v.(n) <- x) | _ -> assert false let prm_array_get : ml_tactic = function -| [ValBlk (_, v); ValInt n] -> wrap (fun () -> v.(n)) +| [ValBlk (_, v); ValInt n] -> + (** FIXME: wrap exception *) + wrap (fun () -> v.(n)) +| _ -> assert false + +(** Int *) + +let prm_int_equal : ml_tactic = function +| [m; n] -> + return (Value.of_bool (Value.to_int m == Value.to_int n)) +| _ -> assert false + +let binop f : ml_tactic = function +| [m; n] -> return (Value.of_int (f (Value.to_int m) (Value.to_int n))) +| _ -> assert false + +let prm_int_compare args = binop Int.compare args +let prm_int_add args = binop (+) args +let prm_int_sub args = binop (-) args +let prm_int_mul args = binop ( * ) args + +let prm_int_neg : ml_tactic = function +| [m] -> return (Value.of_int (~- (Value.to_int m))) +| _ -> assert false + +(** String *) + +let prm_string_make : ml_tactic = function +| [n; c] -> + let n = Value.to_int n in + let c = Value.to_char c in + (** FIXME: wrap exception *) + wrap (fun () -> Value.of_string (Bytes.make n c)) +| _ -> assert false + +let prm_string_length : ml_tactic = function +| [s] -> + return (Value.of_int (Bytes.length (Value.to_string s))) +| _ -> assert false + +let prm_string_set : ml_tactic = function +| [s; n; c] -> + let s = Value.to_string s in + let n = Value.to_int n in + let c = Value.to_char c in + (** FIXME: wrap exception *) + wrap_unit (fun () -> Bytes.set s n c) +| _ -> assert false + +let prm_string_get : ml_tactic = function +| [s; n] -> + let s = Value.to_string s in + let n = Value.to_int n in + (** FIXME: wrap exception *) + wrap (fun () -> Value.of_char (Bytes.get s n)) +| _ -> assert false + +(** Error *) + +let prm_throw : ml_tactic = function +| [e] -> + let (e, info) = Value.to_exn e in + Proofview.tclLIFT (Proofview.NonLogical.raise ~info e) +| _ -> assert false + +(** Control *) + +(** exn -> 'a *) +let prm_zero : ml_tactic = function +| [e] -> + let (e, info) = Value.to_exn e in + Proofview.tclZERO ~info e +| _ -> assert false + +(** exn -> 'a *) +let prm_plus : ml_tactic = function +| [x; k] -> + Proofview.tclOR (thaw x) (fun e -> interp_app k [Value.of_exn e]) +| _ -> assert false + +(** (unit -> 'a) -> 'a *) +let prm_once : ml_tactic = function +| [f] -> Proofview.tclONCE (thaw f) +| _ -> assert false + +(** (unit -> unit) list -> unit *) +let prm_dispatch : ml_tactic = function +| [l] -> + let l = Value.to_list l in + let l = List.map (fun f -> Proofview.tclIGNORE (thaw f)) l in + Proofview.tclDISPATCH l >>= fun () -> return v_unit +| _ -> assert false + +(** (unit -> unit) list -> (unit -> unit) -> (unit -> unit) list -> unit *) +let prm_extend : ml_tactic = function +| [lft; tac; rgt] -> + let lft = Value.to_list lft in + let lft = List.map (fun f -> Proofview.tclIGNORE (thaw f)) lft in + let tac = Proofview.tclIGNORE (thaw tac) in + let rgt = Value.to_list rgt in + let rgt = List.map (fun f -> Proofview.tclIGNORE (thaw f)) rgt in + Proofview.tclEXTEND lft tac rgt >>= fun () -> return v_unit | _ -> assert false +(** (unit -> unit) -> unit *) +let prm_enter : ml_tactic = function +| [f] -> + let f = Proofview.tclIGNORE (thaw f) in + Proofview.tclINDEPENDENT f >>= fun () -> return v_unit +| _ -> assert false + +(** int -> int -> (unit -> 'a) -> 'a *) +let prm_focus : ml_tactic = function +| [i; j; tac] -> + let i = Value.to_int i in + let j = Value.to_int j in + Proofview.tclFOCUS i j (thaw tac) +| _ -> assert false + +(** unit -> unit *) +let prm_shelve : ml_tactic = function +| [_] -> Proofview.shelve >>= fun () -> return v_unit +| _ -> assert false + +(** unit -> unit *) +let prm_shelve_unifiable : ml_tactic = function +| [_] -> Proofview.shelve_unifiable >>= fun () -> return v_unit +| _ -> assert false + +(** unit -> constr *) +let prm_goal : ml_tactic = function +| [_] -> + Proofview.Goal.enter_one { enter = fun gl -> + let concl = Tacmach.New.pf_nf_concl gl in + return (Value.of_constr concl) + } +| _ -> assert false + +(** ident -> constr *) +let prm_hyp : ml_tactic = function +| [id] -> + let id = Value.to_ident id in + pf_apply begin fun env _ -> + let mem = try ignore (Environ.lookup_named id env); true with Not_found -> false in + if mem then return (Value.of_constr (EConstr.mkVar id)) + else Tacticals.New.tclZEROMSG + (str "Hypothesis " ++ quote (Id.print id) ++ str " not found") (** FIXME: Do something more sensible *) + end +| _ -> assert false + +(** (unit -> constr) -> unit *) +let prm_refine : ml_tactic = function +| [c] -> + let c = thaw c >>= fun c -> Proofview.tclUNIT ((), Value.to_constr c) in + Proofview.Goal.nf_enter { enter = fun gl -> + Refine.generic_refine ~unsafe:false c gl + } >>= fun () -> return v_unit +| _ -> assert false + + (** Registering *) let () = Tac2env.define_primitive (pname "print") prm_print let () = Tac2env.define_primitive (pname "message_of_string") prm_message_of_string let () = Tac2env.define_primitive (pname "message_of_int") prm_message_of_int +let () = Tac2env.define_primitive (pname "message_of_constr") prm_message_of_constr +let () = Tac2env.define_primitive (pname "message_concat") prm_message_concat let () = Tac2env.define_primitive (pname "array_make") prm_array_make let () = Tac2env.define_primitive (pname "array_length") prm_array_length let () = Tac2env.define_primitive (pname "array_get") prm_array_get let () = Tac2env.define_primitive (pname "array_set") prm_array_set -(** ML types *) +let () = Tac2env.define_primitive (pname "string_make") prm_string_make +let () = Tac2env.define_primitive (pname "string_length") prm_string_length +let () = Tac2env.define_primitive (pname "string_get") prm_string_get +let () = Tac2env.define_primitive (pname "string_set") prm_string_set + +let () = Tac2env.define_primitive (pname "int_equal") prm_int_equal +let () = Tac2env.define_primitive (pname "int_compare") prm_int_compare +let () = Tac2env.define_primitive (pname "int_neg") prm_int_neg +let () = Tac2env.define_primitive (pname "int_add") prm_int_add +let () = Tac2env.define_primitive (pname "int_sub") prm_int_sub +let () = Tac2env.define_primitive (pname "int_mul") prm_int_mul + +let () = Tac2env.define_primitive (pname "throw") prm_throw + +let () = Tac2env.define_primitive (pname "zero") prm_zero +let () = Tac2env.define_primitive (pname "plus") prm_plus +let () = Tac2env.define_primitive (pname "once") prm_once +let () = Tac2env.define_primitive (pname "dispatch") prm_dispatch +let () = Tac2env.define_primitive (pname "extend") prm_extend +let () = Tac2env.define_primitive (pname "enter") prm_enter + +let () = Tac2env.define_primitive (pname "focus") prm_focus +let () = Tac2env.define_primitive (pname "shelve") prm_shelve +let () = Tac2env.define_primitive (pname "shelve_unifiable") prm_shelve_unifiable +let () = Tac2env.define_primitive (pname "goal") prm_goal +let () = Tac2env.define_primitive (pname "hyp") prm_hyp +let () = Tac2env.define_primitive (pname "refine") prm_refine -let val_tag t = match val_tag t with -| Val.Base t -> t -| _ -> assert false - -let tag_constr = val_tag (topwit Stdarg.wit_constr) +(** ML types *) let constr_flags () = let open Pretyping in @@ -142,28 +419,15 @@ let constr_flags () = expand_evars = true } -(** In Ltac2, the notion of "current environment" only makes sense when there is - at most one goal under focus. Contrarily to Ltac1, instead of dynamically - focussing when we need it, we raise a non-backtracking error when it does - not make sense. *) -exception NonFocussedGoal - -let () = register_handler begin function -| NonFocussedGoal -> str "Several goals under focus" -| _ -> raise Unhandled -end - -let pf_apply f = - Proofview.Goal.goals >>= function - | [] -> - Proofview.tclENV >>= fun env -> - Proofview.tclEVARMAP >>= fun sigma -> - f env sigma - | [gl] -> - gl >>= fun gl -> - f (Proofview.Goal.env gl) (Tacmach.New.project gl) - | _ :: _ :: _ -> - Proofview.tclLIFT (Proofview.NonLogical.raise NonFocussedGoal) +let open_constr_no_classes_flags () = + let open Pretyping in + { + use_typeclasses = false; + solve_unification_constraints = true; + use_hook = Pfedit.solve_by_implicit_tactic (); + fail_evar = false; + expand_evars = true + } (** Embed all Ltac2 data into Values *) let to_lvar ist = @@ -172,17 +436,40 @@ let to_lvar ist = let lfun = Id.Map.map map ist in { empty_lvar with ltac_genargs = lfun } -let () = +let interp_constr flags ist (c, _) = let open Pretyping in - let interp ist (c, _) = pf_apply begin fun env sigma -> + pf_apply begin fun env sigma -> + Proofview.V82.wrap_exceptions begin fun () -> let ist = to_lvar ist in - let (sigma, c) = understand_ltac (constr_flags ()) env sigma ist WithoutTypeConstraint c in - let c = Val.Dyn (tag_constr, c) in + let (sigma, c) = understand_ltac flags env sigma ist WithoutTypeConstraint c in + let c = Val.Dyn (val_constr, c) in Proofview.Unsafe.tclEVARS sigma >>= fun () -> Proofview.tclUNIT c - end in + end + end + +let () = + let open Pretyping in + let interp ist c = interp_constr (constr_flags ()) ist c in let obj = { ml_type = t_constr; ml_interp = interp; } in define_ml_object Stdarg.wit_constr obj + +let () = + let open Pretyping in + let interp ist c = interp_constr (open_constr_no_classes_flags ()) ist c in + let obj = { + ml_type = t_constr; + ml_interp = interp; + } in + define_ml_object Stdarg.wit_open_constr obj + +let () = + let interp _ id = return (Val.Dyn (val_ident, id)) in + let obj = { + ml_type = t_ident; + ml_interp = interp; + } in + define_ml_object Stdarg.wit_ident obj diff --git a/tac2core.mli b/tac2core.mli index 14bde483c1..27144bc6e2 100644 --- a/tac2core.mli +++ b/tac2core.mli @@ -22,13 +22,32 @@ end (** {5 Ltac2 FFI} *) +(** These functions allow to convert back and forth between OCaml and Ltac2 + data representation. The [to_*] functions raise an anomaly whenever the data + has not expected shape. *) + module Value : sig val of_unit : unit -> valexpr val to_unit : valexpr -> unit +val of_int : int -> valexpr +val to_int : valexpr -> int + +val of_bool : bool -> valexpr +val to_bool : valexpr -> bool + +val of_char : char -> valexpr +val to_char : valexpr -> char + val of_list : valexpr list -> valexpr val to_list : valexpr -> valexpr list +val of_constr : EConstr.t -> valexpr +val to_constr : valexpr -> EConstr.t + +val of_exn : Exninfo.iexn -> valexpr +val to_exn : valexpr -> Exninfo.iexn + end diff --git a/tac2entries.ml b/tac2entries.ml index 4098324f12..93ad0ceb0b 100644 --- a/tac2entries.ml +++ b/tac2entries.ml @@ -338,8 +338,7 @@ let register_prim_alg name params def = let def = { typdef_local = false; typdef_expr = def } in ignore (Lib.add_leaf id (inTypDef def)) -let coq_prefix = DirPath.make (List.map Id.of_string ["Ltac2"; "ltac2"; "Coq"]) -let coq_def n = KerName.make2 (MPfile coq_prefix) (Label.make n) +let coq_def n = KerName.make2 Tac2env.coq_prefix (Label.make n) let t_list = coq_def "list" diff --git a/tac2env.ml b/tac2env.ml index bdb8f41ef8..17c70d2e44 100644 --- a/tac2env.ml +++ b/tac2env.ml @@ -203,3 +203,8 @@ module MLType = Genarg.Register(MLTypeObj) let define_ml_object t tpe = MLType.register0 t tpe let interp_ml_object t = MLType.obj t + +(** Absolute paths *) + +let coq_prefix = + MPfile (DirPath.make (List.map Id.of_string ["Init"; "ltac2"; "Coq"])) diff --git a/tac2env.mli b/tac2env.mli index bcfa70487a..96174e8c92 100644 --- a/tac2env.mli +++ b/tac2env.mli @@ -95,3 +95,8 @@ type 'a ml_object = { val define_ml_object : ('a, 'b, 'c) genarg_type -> 'b ml_object -> unit val interp_ml_object : ('a, 'b, 'c) genarg_type -> 'b ml_object + +(** {5 Absolute paths} *) + +val coq_prefix : ModPath.t +(** Path where primitive datatypes are defined in Ltac2 plugin. *) diff --git a/tac2expr.mli b/tac2expr.mli index b9b649e481..c3c1e56dea 100644 --- a/tac2expr.mli +++ b/tac2expr.mli @@ -131,7 +131,7 @@ type valexpr = (** Immediate integers *) | ValBlk of tag * valexpr array (** Structured blocks *) -| ValStr of string +| ValStr of Bytes.t (** Strings *) | ValCls of closure (** Closures *) diff --git a/tac2intern.ml b/tac2intern.ml index 10fcde6efa..23f8325da8 100644 --- a/tac2intern.ml +++ b/tac2intern.ml @@ -18,8 +18,7 @@ open Tac2expr (** Hardwired types and constants *) -let coq_prefix = DirPath.make (List.map Id.of_string ["Ltac2"; "ltac2"; "Coq"]) -let coq_type n = KerName.make2 (MPfile coq_prefix) (Label.make n) +let coq_type n = KerName.make2 Tac2env.coq_prefix (Label.make n) let t_int = coq_type "int" let t_string = coq_type "string" diff --git a/tac2interp.ml b/tac2interp.ml index fedbb13e7d..d508b0c967 100644 --- a/tac2interp.ml +++ b/tac2interp.ml @@ -15,6 +15,8 @@ open Tac2expr exception LtacError of KerName.t * valexpr +let val_exn = Geninterp.Val.create "ltac2:exn" + type environment = valexpr Id.Map.t let empty_environment = Id.Map.empty @@ -35,7 +37,7 @@ let return = Proofview.tclUNIT let rec interp ist = function | GTacAtm (AtmInt n) -> return (ValInt n) -| GTacAtm (AtmStr s) -> return (ValStr s) +| GTacAtm (AtmStr s) -> return (ValStr (Bytes.of_string s)) | GTacVar id -> return (get_var ist id) | GTacRef qid -> return (get_ref ist qid) | GTacFun (ids, e) -> @@ -44,7 +46,7 @@ let rec interp ist = function | GTacApp (f, args) -> interp ist f >>= fun f -> Proofview.Monad.List.map (fun e -> interp ist e) args >>= fun args -> - interp_app ist f args + interp_app f args | GTacLet (false, el, e) -> let fold accu (na, e) = interp ist e >>= fun e -> @@ -94,11 +96,11 @@ let rec interp ist = function let tpe = Tac2env.interp_ml_object tag in tpe.Tac2env.ml_interp ist e >>= fun e -> return (ValExt e) -and interp_app ist f args = match f with +and interp_app f args = match f with | ValCls { clos_env = ist; clos_var = ids; clos_exp = e } -> let rec push ist ids args = match ids, args with | [], [] -> interp ist e - | [], _ :: _ -> interp ist e >>= fun f -> interp_app ist f args + | [], _ :: _ -> interp ist e >>= fun f -> interp_app f args | _ :: _, [] -> let cls = { clos_env = ist; clos_var = ids; clos_exp = e } in return (ValCls cls) diff --git a/tac2interp.mli b/tac2interp.mli index b11ee36012..7fe78a9460 100644 --- a/tac2interp.mli +++ b/tac2interp.mli @@ -10,10 +10,18 @@ open Genarg open Names open Tac2expr -exception LtacError of KerName.t * valexpr - type environment = valexpr Id.Map.t val empty_environment : environment val interp : environment -> glb_tacexpr -> valexpr Proofview.tactic + +val interp_app : valexpr -> valexpr list -> valexpr Proofview.tactic + +(** {5 Exceptions} *) + +exception LtacError of KerName.t * valexpr +(** Ltac2-defined exceptions *) + +val val_exn : Exninfo.iexn Geninterp.Val.typ +(** Toplevel representation of Ltac2 exceptions *) diff --git a/vo.itarget b/vo.itarget index 776404ad79..5777585681 100644 --- a/vo.itarget +++ b/vo.itarget @@ -1 +1,8 @@ +Init.vo +Int.vo +String.vo +Array.vo +Control.vo +Message.vo +Constr.vo Ltac2.vo -- cgit v1.2.3 From d54eacd7b48b9cb0212d5a7cef2ea428469df74a Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 7 Dec 2016 14:44:05 +0100 Subject: Allow the embedding of Ltac2 terms in constrs via the ltac2:(...) syntax. --- g_ltac2.ml4 | 10 ++++++++++ tac2core.ml | 14 ++++++++++++++ tac2env.ml | 4 ++++ tac2env.mli | 4 ++++ tac2intern.ml | 17 +++++++++++++++++ 5 files changed, 49 insertions(+) diff --git a/g_ltac2.ml4 b/g_ltac2.ml4 index ff3d79bbae..1dbd223b22 100644 --- a/g_ltac2.ml4 +++ b/g_ltac2.ml4 @@ -11,6 +11,8 @@ open Util open Genarg open Names open Pcoq +open Constrexpr +open Misctypes open Tac2expr open Ltac_plugin @@ -190,6 +192,14 @@ GEXTEND Gram ; END +GEXTEND Gram + Pcoq.Constr.operconstr: LEVEL "0" + [ [ IDENT "ltac2"; ":"; "("; tac = tac2expr; ")" -> + let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in + CHole (!@loc, None, IntroAnonymous, Some arg) ] ] + ; +END + let pr_ltac2entry _ = mt () (** FIXME *) let pr_ltac2expr _ = mt () (** FIXME *) diff --git a/tac2core.ml b/tac2core.ml index fd998151fd..c18b6032a6 100644 --- a/tac2core.ml +++ b/tac2core.ml @@ -473,3 +473,17 @@ let () = ml_interp = interp; } in define_ml_object Stdarg.wit_ident obj + +let () = + let interp ist env sigma concl tac = + let fold id (Val.Dyn (tag, v)) (accu : environment) : environment = + match Val.eq tag val_valexpr with + | None -> accu + | Some Refl -> Id.Map.add id v accu + in + let ist = Id.Map.fold fold ist Id.Map.empty in + let tac = Proofview.tclIGNORE (interp ist tac) in + let c, sigma = Pfedit.refine_by_tactic env sigma concl tac in + (EConstr.of_constr c, sigma) + in + Pretyping.register_constr_interp0 wit_ltac2 interp diff --git a/tac2env.ml b/tac2env.ml index 17c70d2e44..a058d764e7 100644 --- a/tac2env.ml +++ b/tac2env.ml @@ -208,3 +208,7 @@ let interp_ml_object t = MLType.obj t let coq_prefix = MPfile (DirPath.make (List.map Id.of_string ["Init"; "ltac2"; "Coq"])) + +(** Generic arguments *) + +let wit_ltac2 = Genarg.make0 "ltac2" diff --git a/tac2env.mli b/tac2env.mli index 96174e8c92..4d2a1645ea 100644 --- a/tac2env.mli +++ b/tac2env.mli @@ -100,3 +100,7 @@ val interp_ml_object : ('a, 'b, 'c) genarg_type -> 'b ml_object val coq_prefix : ModPath.t (** Path where primitive datatypes are defined in Ltac2 plugin. *) + +(** {5 Generic arguments} *) + +val wit_ltac2 : (raw_tacexpr, glb_tacexpr, Util.Empty.t) genarg_type diff --git a/tac2intern.ml b/tac2intern.ml index 23f8325da8..bc15b567d4 100644 --- a/tac2intern.ml +++ b/tac2intern.ml @@ -1054,3 +1054,20 @@ let subst_quant_typedef subst (prm, def as qdef) = let subst_type_scheme subst (prm, t as sch) = let t' = subst_type subst t in if t' == t then sch else (prm, t') + +(** Registering *) + +let () = + let open Genintern in + let intern ist tac = + let env = match Genintern.Store.get ist.extra ltac2_env with + | None -> empty_env () + | Some env -> env + in + let loc = loc_of_tacexpr tac in + let (tac, t) = intern_rec env tac in + let () = check_elt_unit loc env t in + (ist, tac) + in + Genintern.register_intern0 wit_ltac2 intern +let () = Genintern.register_subst0 wit_ltac2 subst_expr -- cgit v1.2.3 From f0b3169d5494074d159f94ed1d3d482037990a58 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 7 Dec 2016 18:14:23 +0100 Subject: Towards a proper printing of Ltac2 data structures. --- g_ltac2.ml4 | 2 +- ltac2_plugin.mlpack | 1 + tac2entries.ml | 13 ++- tac2env.ml | 42 ++++++++- tac2env.mli | 5 +- tac2expr.mli | 4 +- tac2intern.ml | 69 +++++++++----- tac2intern.mli | 2 - tac2interp.ml | 7 +- tac2print.ml | 266 ++++++++++++++++++++++++++++++++++++++++++++++++++++ tac2print.mli | 28 ++++++ 11 files changed, 401 insertions(+), 38 deletions(-) create mode 100644 tac2print.ml create mode 100644 tac2print.mli diff --git a/g_ltac2.ml4 b/g_ltac2.ml4 index 1dbd223b22..9384584f19 100644 --- a/g_ltac2.ml4 +++ b/g_ltac2.ml4 @@ -56,7 +56,7 @@ GEXTEND Gram [ e = tac2expr; el = LIST1 tac2expr LEVEL "0" -> CTacApp (!@loc, e, el) | e = SELF; ".("; qid = Prim.qualid; ")" -> CTacPrj (!@loc, e, qid) | e = SELF; ".("; qid = Prim.qualid; ")"; ":="; r = tac2expr LEVEL "1" -> CTacSet (!@loc, e, qid, r) - | e0 = tac2expr; ","; el = LIST1 tac2expr LEVEL "0" SEP "," -> CTacTup (!@loc, e0 :: el) ] + | e0 = tac2expr; ","; el = LIST1 tac2expr LEVEL "1" SEP "," -> CTacTup (!@loc, e0 :: el) ] | "0" [ "("; a = tac2expr LEVEL "5"; ")" -> a | "("; a = tac2expr; ":"; t = tac2type; ")" -> CTacCnv (!@loc, a, t) diff --git a/ltac2_plugin.mlpack b/ltac2_plugin.mlpack index 561bd0eb0a..3d87a8cddb 100644 --- a/ltac2_plugin.mlpack +++ b/ltac2_plugin.mlpack @@ -1,4 +1,5 @@ Tac2env +Tac2print Tac2intern Tac2interp Tac2entries diff --git a/tac2entries.ml b/tac2entries.ml index 93ad0ceb0b..f4b3147c48 100644 --- a/tac2entries.ml +++ b/tac2entries.ml @@ -15,6 +15,7 @@ open Libobject open Nametab open Tac2env open Tac2expr +open Tac2print open Tac2intern open Vernacexpr @@ -303,11 +304,17 @@ let print_ltac ref = in match kn with | TacConstant kn -> - let (_, (_, t)) = Tac2env.interp_global kn in - Feedback.msg_notice (pr_qualid qid ++ spc () ++ str ":" ++ spc () ++ pr_glbtype t) + let (e, _, (_, t)) = Tac2env.interp_global kn in + let name = int_name () in + Feedback.msg_notice ( + hov 0 ( + hov 2 (pr_qualid qid ++ spc () ++ str ":" ++ spc () ++ pr_glbtype name t) ++ fnl () ++ + hov 2 (pr_qualid qid ++ spc () ++ str ":=" ++ spc () ++ pr_glbexpr e) + ) + ) | TacConstructor kn -> let _ = Tac2env.interp_constructor kn in - Feedback.msg_notice (str "Constructor" ++ spc () ++ str ":" ++ spc () ++ pr_qualid qid) + Feedback.msg_notice (hov 2 (str "Constructor" ++ spc () ++ str ":" ++ spc () ++ pr_qualid qid)) (** Calling tactics *) diff --git a/tac2env.ml b/tac2env.ml index a058d764e7..18519a6ce1 100644 --- a/tac2env.ml +++ b/tac2env.ml @@ -73,7 +73,7 @@ let define_global kn e = let interp_global kn = let (e, t) = KNmap.find kn ltac_state.contents.ltac_tactics in - (eval_pure e, t) + (e, eval_pure e, t) let define_constructor kn t = let state = !ltac_state in @@ -139,21 +139,33 @@ module RfTab = Nametab.Make(FullPath)(TacRef) type nametab = { tab_ltac : RfTab.t; + tab_ltac_rev : full_path KNmap.t * full_path KNmap.t; tab_type : KnTab.t; + tab_type_rev : full_path KNmap.t; tab_proj : KnTab.t; + tab_proj_rev : full_path KNmap.t; } let empty_nametab = { tab_ltac = RfTab.empty; + tab_ltac_rev = (KNmap.empty, KNmap.empty); tab_type = KnTab.empty; + tab_type_rev = KNmap.empty; tab_proj = KnTab.empty; + tab_proj_rev = KNmap.empty; } let nametab = Summary.ref empty_nametab ~name:"ltac2-nametab" let push_ltac vis sp kn = let tab = !nametab in - nametab := { tab with tab_ltac = RfTab.push vis sp kn tab.tab_ltac } + let tab_ltac = RfTab.push vis sp kn tab.tab_ltac in + let (constant_map, constructor_map) = tab.tab_ltac_rev in + let tab_ltac_rev = match kn with + | TacConstant c -> (KNmap.add c sp constant_map, constructor_map) + | TacConstructor c -> (constant_map, KNmap.add c sp constructor_map) + in + nametab := { tab with tab_ltac; tab_ltac_rev } let locate_ltac qid = let tab = !nametab in @@ -163,9 +175,19 @@ let locate_extended_all_ltac qid = let tab = !nametab in RfTab.find_prefixes qid tab.tab_ltac +let shortest_qualid_of_ltac kn = + let tab = !nametab in + let sp = match kn with + | TacConstant c -> KNmap.find c (fst tab.tab_ltac_rev) + | TacConstructor c -> KNmap.find c (snd tab.tab_ltac_rev) + in + RfTab.shortest_qualid Id.Set.empty sp tab.tab_ltac + let push_type vis sp kn = let tab = !nametab in - nametab := { tab with tab_type = KnTab.push vis sp kn tab.tab_type } + let tab_type = KnTab.push vis sp kn tab.tab_type in + let tab_type_rev = KNmap.add kn sp tab.tab_type_rev in + nametab := { tab with tab_type; tab_type_rev } let locate_type qid = let tab = !nametab in @@ -175,9 +197,16 @@ let locate_extended_all_type qid = let tab = !nametab in KnTab.find_prefixes qid tab.tab_type +let shortest_qualid_of_type kn = + let tab = !nametab in + let sp = KNmap.find kn tab.tab_type_rev in + KnTab.shortest_qualid Id.Set.empty sp tab.tab_type + let push_projection vis sp kn = let tab = !nametab in - nametab := { tab with tab_proj = KnTab.push vis sp kn tab.tab_proj } + let tab_proj = KnTab.push vis sp kn tab.tab_proj in + let tab_proj_rev = KNmap.add kn sp tab.tab_proj_rev in + nametab := { tab with tab_proj; tab_proj_rev } let locate_projection qid = let tab = !nametab in @@ -187,6 +216,11 @@ let locate_extended_all_projection qid = let tab = !nametab in KnTab.find_prefixes qid tab.tab_proj +let shortest_qualid_of_projection kn = + let tab = !nametab in + let sp = KNmap.find kn tab.tab_proj_rev in + KnTab.shortest_qualid Id.Set.empty sp tab.tab_proj + type 'a ml_object = { ml_type : type_constant; ml_interp : environment -> 'a -> Geninterp.Val.t Proofview.tactic; diff --git a/tac2env.mli b/tac2env.mli index 4d2a1645ea..16232ec810 100644 --- a/tac2env.mli +++ b/tac2env.mli @@ -17,7 +17,7 @@ open Tac2expr (** {5 Toplevel definition of values} *) val define_global : ltac_constant -> (glb_tacexpr * type_scheme) -> unit -val interp_global : ltac_constant -> (valexpr * type_scheme) +val interp_global : ltac_constant -> (glb_tacexpr * valexpr * type_scheme) (** {5 Toplevel definition of types} *) @@ -69,14 +69,17 @@ type tacref = val push_ltac : visibility -> full_path -> tacref -> unit val locate_ltac : qualid -> tacref val locate_extended_all_ltac : qualid -> tacref list +val shortest_qualid_of_ltac : tacref -> qualid val push_type : visibility -> full_path -> type_constant -> unit val locate_type : qualid -> type_constant val locate_extended_all_type : qualid -> type_constant list +val shortest_qualid_of_type : type_constant -> qualid val push_projection : visibility -> full_path -> ltac_projection -> unit val locate_projection : qualid -> ltac_projection val locate_extended_all_projection : qualid -> ltac_projection list +val shortest_qualid_of_projection : ltac_projection -> qualid (** {5 Toplevel definitions of ML tactics} *) diff --git a/tac2expr.mli b/tac2expr.mli index c3c1e56dea..7a2c684fbc 100644 --- a/tac2expr.mli +++ b/tac2expr.mli @@ -105,8 +105,8 @@ type glb_tacexpr = | GTacArr of glb_tacexpr list | GTacCst of type_constant * int * glb_tacexpr list | GTacCse of glb_tacexpr * case_info * glb_tacexpr array * (Name.t array * glb_tacexpr) array -| GTacPrj of glb_tacexpr * int -| GTacSet of glb_tacexpr * int * glb_tacexpr +| GTacPrj of type_constant * glb_tacexpr * int +| GTacSet of type_constant * glb_tacexpr * int * glb_tacexpr | GTacExt of glob_generic_argument | GTacPrm of ml_tactic_name * glb_tacexpr list diff --git a/tac2intern.ml b/tac2intern.ml index bc15b567d4..350dc4efe6 100644 --- a/tac2intern.ml +++ b/tac2intern.ml @@ -14,6 +14,7 @@ open Names open Libnames open Misctypes open Tac2env +open Tac2print open Tac2expr (** Hardwired types and constants *) @@ -50,6 +51,7 @@ sig val add : key -> 'a -> 'a t -> 'a t val mem : key -> 'a t -> bool val find : key -> 'a t -> 'a + val exists : (key -> 'a -> bool) -> 'a t -> bool end end = @@ -143,6 +145,31 @@ let empty_env () = { env_rec = Id.Map.empty; } +let env_name env = + (** Generate names according to a provided environment *) + let mk num = + let base = num mod 26 in + let rem = num / 26 in + let name = String.make 1 (Char.chr (97 + base)) in + let suff = if Int.equal rem 0 then "" else string_of_int rem in + let name = name ^ suff in + name + in + let fold id elt acc = UF.Map.add elt (Id.to_string id) acc in + let vars = Id.Map.fold fold env.env_als.contents UF.Map.empty in + let vars = ref vars in + let rec fresh n = + let name = mk n in + if UF.Map.exists (fun _ name' -> String.equal name name') !vars then fresh (succ n) + else name + in + fun n -> + if UF.Map.mem n !vars then UF.Map.find n !vars + else + let ans = fresh 0 in + let () = vars := UF.Map.add n ans !vars in + ans + let ltac2_env : environment Genintern.Store.field = Genintern.Store.field () @@ -311,19 +338,12 @@ let rec unify env t1 t2 = match kind env t1, kind env t2 with else raise (CannotUnify (t1, t2)) | _ -> raise (CannotUnify (t1, t2)) -(** FIXME *) -let rec pr_glbtype = function -| GTypVar n -> str "?" -| GTypRef (kn, tl) -> - KerName.print kn ++ str "(" ++ prlist_with_sep (fun () -> str ", ") pr_glbtype tl ++ str ")" -| GTypArrow (t1, t2) -> str "Arr(" ++ pr_glbtype t1 ++ str ", " ++ pr_glbtype t2 ++ str ")" -| GTypTuple tl -> str "Tup(" ++ prlist_with_sep (fun () -> str ", ") pr_glbtype tl ++ str ")" - let unify loc env t1 t2 = try unify env t1 t2 with CannotUnify (u1, u2) -> - user_err ~loc (str "This expression has type " ++ pr_glbtype t1 ++ - str " but an expression what expected of type " ++ pr_glbtype t2) + let name = env_name env in + user_err ~loc (str "This expression has type " ++ pr_glbtype name t1 ++ + str " but an expression what expected of type " ++ pr_glbtype name t2) (** Term typing *) @@ -418,13 +438,15 @@ let check_elt_empty loc env t = match kind env t with | GTypVar _ -> user_err ~loc (str "Cannot infer an empty type for this expression") | GTypArrow _ | GTypTuple _ -> - user_err ~loc (str "Type " ++ pr_glbtype t ++ str " is not an empty type") + let name = env_name env in + user_err ~loc (str "Type " ++ pr_glbtype name t ++ str " is not an empty type") | GTypRef (kn, _) -> let def = Tac2env.interp_type kn in match def with | _, GTydAlg [] -> kn | _ -> - user_err ~loc (str "Type " ++ pr_glbtype t ++ str " is not an empty type") + let name = env_name env in + user_err ~loc (str "Type " ++ pr_glbtype name t ++ str " is not an empty type") let check_unit ?loc t = let maybe_unit = match t with @@ -475,11 +497,11 @@ let intern_atm env = function let invalid_pattern ~loc kn t = let pt = match t with - | GCaseAlg kn' -> KerName.print kn + | GCaseAlg kn' -> pr_typref kn | GCaseTuple n -> str "tuple" in user_err ~loc (str "Invalid pattern, expected a pattern for type " ++ - KerName.print kn ++ str ", found a pattern of type " ++ pt) (** FIXME *) + pr_typref kn ++ str ", found a pattern of type " ++ pt) (** FIXME *) type pattern_kind = | PKind_empty @@ -527,7 +549,7 @@ let rec intern_rec env = function let sch = Id.Map.find id env.env_var in (GTacVar id, fresh_mix_type_scheme env sch) | ArgArg (TacConstant kn) -> - let (_, sch) = Tac2env.interp_global kn in + let (_, _, sch) = Tac2env.interp_global kn in (GTacRef kn, fresh_type_scheme env sch) | ArgArg (TacConstructor kn) -> intern_constructor env (fst qid) kn [] @@ -542,6 +564,7 @@ let rec intern_rec env = function (env, na :: bnd, t :: tl) in let (env, bnd, tl) = List.fold_left fold (env, [], []) bnd in + let bnd = List.rev bnd in let (e, t) = intern_rec env e in let t = List.fold_left (fun accu t -> GTypArrow (t, accu)) t tl in (GTacFun (bnd, e), t) @@ -637,7 +660,7 @@ let rec intern_rec env = function let () = unify loc env t exp in let substf i = GTypVar subst.(i) in let ret = subst_type substf pinfo.pdata_ptyp in - (GTacPrj (e, pinfo.pdata_indx), ret) + (GTacPrj (pinfo.pdata_type, e, pinfo.pdata_indx), ret) | CTacSet (loc, e, proj, r) -> let pinfo = get_projection proj in let () = @@ -652,7 +675,7 @@ let rec intern_rec env = function let substf i = GTypVar subst.(i) in let ret = subst_type substf pinfo.pdata_ptyp in let r = intern_rec_with_constraint env r ret in - (GTacSet (e, pinfo.pdata_indx, r), GTypRef (t_unit, [])) + (GTacSet (pinfo.pdata_type, e, pinfo.pdata_indx, r), GTypRef (t_unit, [])) | CTacExt (loc, ext) -> let open Genintern in let GenArg (Rawwit tag, _) = ext in @@ -899,7 +922,7 @@ and intern_record env loc fs = several times") else user_err ~loc (str "Field " ++ (*KerName.print knp ++*) str " does not \ - pertain to record definition " ++ KerName.print pinfo.pdata_type) + pertain to record definition " ++ pr_typref pinfo.pdata_type) in let () = List.iter iter fs in let () = match Array.findi (fun _ o -> Option.is_empty o) args with @@ -1017,13 +1040,15 @@ let rec subst_expr subst e = match e with let cse1' = Array.map (fun (ids, e) -> (ids, subst_expr subst e)) cse1 in let ci' = subst_case_info subst ci in GTacCse (subst_expr subst e, ci', cse0', cse1') -| GTacPrj (e, p) as e0 -> +| GTacPrj (kn, e, p) as e0 -> + let kn' = subst_kn subst kn in let e' = subst_expr subst e in - if e' == e then e0 else GTacPrj (e', p) -| GTacSet (e, p, r) as e0 -> + if kn' == kn && e' == e then e0 else GTacPrj (kn', e', p) +| GTacSet (kn, e, p, r) as e0 -> + let kn' = subst_kn subst kn in let e' = subst_expr subst e in let r' = subst_expr subst r in - if e' == e && r' == r then e0 else GTacSet (e', p, r') + if kn' == kn && e' == e && r' == r then e0 else GTacSet (kn', e', p, r') | GTacExt ext -> let ext' = Genintern.generic_substitute subst ext in if ext' == ext then e else GTacExt ext' diff --git a/tac2intern.mli b/tac2intern.mli index a6be01d647..6633792e7e 100644 --- a/tac2intern.mli +++ b/tac2intern.mli @@ -26,5 +26,3 @@ val subst_type : substitution -> 'a glb_typexpr -> 'a glb_typexpr val subst_expr : substitution -> glb_tacexpr -> glb_tacexpr val subst_quant_typedef : substitution -> glb_quant_typedef -> glb_quant_typedef val subst_type_scheme : substitution -> type_scheme -> type_scheme - -val pr_glbtype : 'a glb_typexpr -> Pp.std_ppcmds diff --git a/tac2interp.ml b/tac2interp.ml index d508b0c967..f93c8cb5fe 100644 --- a/tac2interp.ml +++ b/tac2interp.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open Util open Pp open CErrors open Genarg @@ -30,7 +31,7 @@ let get_var ist id = anomaly (str "Unbound variable " ++ Id.print id) let get_ref ist kn = - try fst (Tac2env.interp_global kn) with Not_found -> + try pi2 (Tac2env.interp_global kn) with Not_found -> anomaly (str "Unbound reference" ++ KerName.print kn) let return = Proofview.tclUNIT @@ -82,9 +83,9 @@ let rec interp ist = function return (ValBlk (n, Array.of_list el)) | GTacCse (e, _, cse0, cse1) -> interp ist e >>= fun e -> interp_case ist e cse0 cse1 -| GTacPrj (e, p) -> +| GTacPrj (_, e, p) -> interp ist e >>= fun e -> interp_proj ist e p -| GTacSet (e, p, r) -> +| GTacSet (_, e, p, r) -> interp ist e >>= fun e -> interp ist r >>= fun r -> interp_set ist e p r diff --git a/tac2print.ml b/tac2print.ml new file mode 100644 index 0000000000..96d0ceb875 --- /dev/null +++ b/tac2print.ml @@ -0,0 +1,266 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* str "'" ++ str (pr n) + | GTypRef (kn, []) -> pr_typref kn + | GTypRef (kn, [t]) -> + let paren = match lvl with + | T5_r | T5_l | T2 | T1 -> fun x -> x + | T0 -> paren + in + paren (pr_glbtype lvl t ++ spc () ++ pr_typref kn) + | GTypRef (kn, tl) -> + let paren = match lvl with + | T5_r | T5_l | T2 | T1 -> fun x -> x + | T0 -> paren + in + paren (str "(" ++ prlist_with_sep (fun () -> str ", ") (pr_glbtype lvl) tl ++ str ")" ++ spc () ++ pr_typref kn) + | GTypArrow (t1, t2) -> + let paren = match lvl with + | T5_r -> fun x -> x + | T5_l | T2 | T1 | T0 -> paren + in + paren (pr_glbtype T5_l t1 ++ spc () ++ str "->" ++ spc () ++ pr_glbtype T5_r t2) + | GTypTuple tl -> + let paren = match lvl with + | T5_r | T5_l -> fun x -> x + | T2 | T1 | T0 -> paren + in + paren (prlist_with_sep (fun () -> str " * ") (pr_glbtype T2) tl) + in + hov 0 (pr_glbtype lvl c) + +let pr_glbtype pr c = pr_glbtype_gen pr T5_r c + +let int_name () = + let vars = ref Int.Map.empty in + fun n -> + if Int.Map.mem n !vars then Int.Map.find n !vars + else + let num = Int.Map.cardinal !vars in + let base = num mod 26 in + let rem = num / 26 in + let name = String.make 1 (Char.chr (97 + base)) in + let suff = if Int.equal rem 0 then "" else string_of_int rem in + let name = name ^ suff in + let () = vars := Int.Map.add n name !vars in + name + +(** Term printing *) + +let pr_constructor kn = + Libnames.pr_qualid (Tac2env.shortest_qualid_of_ltac (TacConstructor kn)) + +let pr_projection kn = + Libnames.pr_qualid (Tac2env.shortest_qualid_of_projection kn) + +type exp_level = +| E5 +| E4 +| E3 +| E2 +| E1 +| E0 + +let pr_atom = function +| AtmInt n -> int n +| AtmStr s -> qstring s + +let pr_name = function +| Name id -> Id.print id +| Anonymous -> str "_" + +let find_constructor n empty def = + let rec find n = function + | [] -> assert false + | (id, []) :: rem -> + if empty then + if Int.equal n 0 then id + else find (pred n) rem + else find n rem + | (id, _ :: _) :: rem -> + if not empty then + if Int.equal n 0 then id + else find (pred n) rem + else find n rem + in + find n def + +let order_branches cbr nbr def = + let rec order cidx nidx def = match def with + | [] -> [] + | (id, []) :: rem -> + let ans = order (succ cidx) nidx rem in + (id, [], cbr.(cidx)) :: ans + | (id, _ :: _) :: rem -> + let ans = order cidx (succ nidx) rem in + let (vars, e) = nbr.(nidx) in + (id, Array.to_list vars, e) :: ans + in + order 0 0 def + +let pr_glbexpr_gen lvl c = + let rec pr_glbexpr lvl = function + | GTacAtm atm -> pr_atom atm + | GTacVar id -> Id.print id + | GTacRef gr -> + let qid = shortest_qualid_of_ltac (TacConstant gr) in + Libnames.pr_qualid qid + | GTacFun (nas, c) -> + let nas = pr_sequence pr_name nas in + let paren = match lvl with + | E0 | E1 | E2 | E3 | E4 -> paren + | E5 -> fun x -> x + in + paren (str "fun" ++ spc () ++ nas ++ spc () ++ str "=>" ++ spc () ++ + hov 0 (pr_glbexpr E5 c)) + | GTacApp (c, cl) -> + let paren = match lvl with + | E0 -> paren + | E1 | E2 | E3 | E4 | E5 -> fun x -> x + in + paren (pr_glbexpr E1 c) ++ spc () ++ (pr_sequence (pr_glbexpr E0) cl) + | GTacLet (mut, bnd, e) -> + let paren = match lvl with + | E0 | E1 | E2 | E3 | E4 -> paren + | E5 -> fun x -> x + in + let mut = if mut then str "rec" ++ spc () else mt () in + let pr_bnd (na, e) = + pr_name na ++ spc () ++ str ":=" ++ spc () ++ hov 2 (pr_glbexpr E5 e) ++ spc () + in + let bnd = prlist_with_sep (fun () -> str "with" ++ spc ()) pr_bnd bnd in + paren (str "let" ++ spc () ++ mut ++ bnd ++ str "in" ++ spc () ++ pr_glbexpr E5 e) + | GTacTup cl -> + let paren = match lvl with + | E0 | E1 -> paren + | E2 | E3 | E4 | E5 -> fun x -> x + in + paren (prlist_with_sep (fun () -> str "," ++ spc ()) (pr_glbexpr E1) cl) + | GTacArr cl -> + mt () (** FIXME when implemented *) + | GTacCst (tpe, n, cl) -> + begin match Tac2env.interp_type tpe with + | _, GTydAlg def -> + let paren = match lvl with + | E0 -> paren + | E1 | E2 | E3 | E4 | E5 -> fun x -> x + in + let id = find_constructor n (List.is_empty cl) def in + let kn = change_kn_label tpe id in + let cl = match cl with + | [] -> mt () + | _ -> spc () ++ pr_sequence (pr_glbexpr E0) cl + in + paren (pr_constructor kn ++ cl) + | _, GTydRec def -> + let args = List.combine def cl in + let pr_arg ((id, _, _), arg) = + let kn = change_kn_label tpe id in + pr_projection kn ++ spc () ++ str ":=" ++ spc () ++ pr_glbexpr E1 arg + in + let args = prlist_with_sep (fun () -> str ";" ++ spc ()) pr_arg args in + str "{" ++ spc () ++ args ++ spc () ++ str "}" + | _, GTydDef _ -> assert false + end + | GTacCse (e, info, cst_br, ncst_br) -> + let e = pr_glbexpr E5 e in + let br = match info with + | GCaseAlg kn -> + let def = match Tac2env.interp_type kn with + | _, GTydAlg def -> def + | _, GTydDef _ | _, GTydRec _ -> assert false + in + let br = order_branches cst_br ncst_br def in + let pr_branch (cstr, vars, p) = + let cstr = change_kn_label kn cstr in + let cstr = pr_constructor cstr in + let vars = match vars with + | [] -> mt () + | _ -> spc () ++ pr_sequence pr_name vars + in + hov 0 (str "|" ++ spc () ++ cstr ++ vars ++ spc () ++ str "=>" ++ spc () ++ + hov 2 (pr_glbexpr E5 p)) ++ spc () + in + prlist pr_branch br + | GCaseTuple n -> + let (vars, p) = ncst_br.(0) in + let p = pr_glbexpr E5 p in + let vars = prvect_with_sep (fun () -> str "," ++ spc ()) pr_name vars in + str "|" ++ spc () ++ paren vars ++ spc () ++ str "=>" ++ spc () ++ p + in + hov 0 (hov 0 (str "match" ++ spc () ++ e ++ spc () ++ str "with") ++ spc () ++ Pp.v 0 br ++ str "end") + | GTacPrj (kn, e, n) -> + let def = match Tac2env.interp_type kn with + | _, GTydRec def -> def + | _, GTydDef _ | _, GTydAlg _ -> assert false + in + let (proj, _, _) = List.nth def n in + let proj = change_kn_label kn proj in + let proj = pr_projection proj in + let e = pr_glbexpr E0 e in + e ++ str "." ++ paren proj + | GTacSet (kn, e, n, r) -> + let def = match Tac2env.interp_type kn with + | _, GTydRec def -> def + | _, GTydDef _ | _, GTydAlg _ -> assert false + in + let (proj, _, _) = List.nth def n in + let proj = change_kn_label kn proj in + let proj = pr_projection proj in + let e = pr_glbexpr E0 e in + let r = pr_glbexpr E1 r in + e ++ str "." ++ paren proj ++ spc () ++ str ":=" ++ spc () ++ r + | GTacExt arg -> + let GenArg (Glbwit tag, arg) = arg in + let name = match tag with + | ExtraArg tag -> ArgT.repr tag + | _ -> assert false + in + str name ++ str ":" ++ paren (Genprint.glb_print tag arg) + | GTacPrm (prm, args) -> + let args = match args with + | [] -> mt () + | _ -> spc () ++ pr_sequence (pr_glbexpr E0) args + in + str "@external" ++ spc () ++ qstring prm.mltac_plugin ++ spc () ++ + qstring prm.mltac_tactic ++ args + in + hov 0 (pr_glbexpr lvl c) + +let pr_glbexpr c = + pr_glbexpr_gen E5 c diff --git a/tac2print.mli b/tac2print.mli new file mode 100644 index 0000000000..94555a4c95 --- /dev/null +++ b/tac2print.mli @@ -0,0 +1,28 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* std_ppcmds +val pr_glbtype : ('a -> string) -> 'a glb_typexpr -> std_ppcmds + +(** {5 Printing expressions} *) + +val pr_constructor : ltac_constructor -> std_ppcmds +val pr_projection : ltac_projection -> std_ppcmds +val pr_glbexpr : glb_tacexpr -> std_ppcmds + +(** {5 Utilities} *) + +val int_name : unit -> (int -> string) +(** Create a function that give names to integers. The names are generated on + the fly, in the order they are encountered. *) -- cgit v1.2.3 From 4c5f635769811be7d5f8b39f699b76ea51388cd4 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 8 Dec 2016 09:10:39 +0100 Subject: Merging GTacTuple and GTacCst nodes. --- tac2env.ml | 1 - tac2expr.mli | 3 +-- tac2intern.ml | 31 ++++++++++++++++++------------- tac2interp.ml | 4 +--- tac2print.ml | 4 ++-- 5 files changed, 22 insertions(+), 21 deletions(-) diff --git a/tac2env.ml b/tac2env.ml index 18519a6ce1..5e379473c8 100644 --- a/tac2env.ml +++ b/tac2env.ml @@ -60,7 +60,6 @@ let rec eval_pure = function eval_pure e | GTacFun (na, e) -> ValCls { clos_env = Id.Map.empty; clos_var = na; clos_exp = e } -| GTacTup el -> ValBlk (0, Array.map_of_list eval_pure el) | GTacCst (_, n, []) -> ValInt n | GTacCst (_, n, el) -> ValBlk (n, Array.map_of_list eval_pure el) | GTacAtm (AtmStr _) | GTacArr _ | GTacLet _ | GTacVar _ | GTacSet _ diff --git a/tac2expr.mli b/tac2expr.mli index 7a2c684fbc..1fac5a2315 100644 --- a/tac2expr.mli +++ b/tac2expr.mli @@ -101,9 +101,8 @@ type glb_tacexpr = | GTacFun of Name.t list * glb_tacexpr | GTacApp of glb_tacexpr * glb_tacexpr list | GTacLet of rec_flag * (Name.t * glb_tacexpr) list * glb_tacexpr -| GTacTup of glb_tacexpr list | GTacArr of glb_tacexpr list -| GTacCst of type_constant * int * glb_tacexpr list +| GTacCst of case_info * int * glb_tacexpr list | GTacCse of glb_tacexpr * case_info * glb_tacexpr array * (Name.t array * glb_tacexpr) array | GTacPrj of type_constant * glb_tacexpr * int | GTacSet of type_constant * glb_tacexpr * int * glb_tacexpr diff --git a/tac2intern.ml b/tac2intern.ml index 350dc4efe6..c4d2fc277b 100644 --- a/tac2intern.ml +++ b/tac2intern.ml @@ -27,8 +27,8 @@ let t_array = coq_type "array" let t_unit = coq_type "unit" let t_list = coq_type "list" -let c_nil = GTacCst (t_list, 0, []) -let c_cons e el = GTacCst (t_list, 0, [e; el]) +let c_nil = GTacCst (GCaseAlg t_list, 0, []) +let c_cons e el = GTacCst (GCaseAlg t_list, 0, [e; el]) (** Union find *) @@ -358,15 +358,15 @@ let is_pure_constructor kn = let rec is_value = function | GTacAtm (AtmInt _) | GTacVar _ | GTacRef _ | GTacFun _ -> true | GTacAtm (AtmStr _) | GTacApp _ | GTacLet _ -> false -| GTacTup el -> List.for_all is_value el +| GTacCst (GCaseTuple _, _, el) -> List.for_all is_value el | GTacCst (_, _, []) -> true -| GTacCst (kn, _, el) -> is_pure_constructor kn && List.for_all is_value el +| GTacCst (GCaseAlg kn, _, el) -> is_pure_constructor kn && List.for_all is_value el | GTacArr _ | GTacCse _ | GTacPrj _ | GTacSet _ | GTacExt _ | GTacPrm _ -> false let is_rec_rhs = function | GTacFun _ -> true | GTacAtm _ | GTacVar _ | GTacRef _ | GTacApp _ | GTacLet _ | GTacPrj _ -| GTacSet _ | GTacTup _ | GTacArr _ | GTacExt _ | GTacPrm _ | GTacCst _ +| GTacSet _ | GTacArr _ | GTacExt _ | GTacPrm _ | GTacCst _ | GTacCse _ -> false let rec fv_type f t accu = match t with @@ -612,14 +612,14 @@ let rec intern_rec env = function | CTacLet (loc, true, el, e) -> intern_let_rec env loc el e | CTacTup (loc, []) -> - (GTacTup [], GTypRef (t_unit, [])) + (GTacCst (GCaseAlg t_unit, 0, []), GTypRef (t_unit, [])) | CTacTup (loc, el) -> let fold e (el, tl) = let (e, t) = intern_rec env e in (e :: el, t :: tl) in let (el, tl) = List.fold_right fold el ([], []) in - (GTacTup el, GTypTuple tl) + (GTacCst (GCaseTuple (List.length el), 0, el), GTypTuple tl) | CTacArr (loc, []) -> let id = fresh_id env in (GTacArr [], GTypRef (t_int, [GTypVar id])) @@ -887,7 +887,7 @@ and intern_constructor env loc kn args = let ans = GTypRef (cstr.cdata_type, List.init cstr.cdata_prms (fun i -> GTypVar subst.(i))) in let map arg tpe = intern_rec_with_constraint env arg tpe in let args = List.map2 map args types in - (GTacCst (cstr.cdata_type, cstr.cdata_indx, args), ans) + (GTacCst (GCaseAlg cstr.cdata_type, cstr.cdata_indx, args), ans) else error_nargs_mismatch loc nargs (List.length args) @@ -933,7 +933,7 @@ and intern_record env loc fs = in let args = Array.map_to_list Option.get args in let tparam = List.init params (fun i -> GTypVar subst.(i)) in - (GTacCst (kn, 0, args), GTypRef (kn, tparam)) + (GTacCst (GCaseAlg kn, 0, args), GTypRef (kn, tparam)) let normalize env (count, vars) (t : UF.elt glb_typexpr) = let get_var id = @@ -1029,12 +1029,17 @@ let rec subst_expr subst e = match e with | GTacLet (r, bs, e) -> let bs = List.map (fun (na, e) -> (na, subst_expr subst e)) bs in GTacLet (r, bs, subst_expr subst e) -| GTacTup el -> - GTacTup (List.map (fun e -> subst_expr subst e) el) | GTacArr el -> GTacArr (List.map (fun e -> subst_expr subst e) el) -| GTacCst (kn, n, el) -> - GTacCst (subst_kn subst kn, n, List.map (fun e -> subst_expr subst e) el) +| GTacCst (t, n, el) as e0 -> + let t' = match t with + | GCaseAlg kn -> + let kn' = subst_kn subst kn in + if kn' == kn then t else GCaseAlg kn' + | GCaseTuple _ -> t + in + let el' = List.smartmap (fun e -> subst_expr subst e) el in + if t' == t && el' == el then e0 else GTacCst (t', n, el') | GTacCse (e, ci, cse0, cse1) -> let cse0' = Array.map (fun e -> subst_expr subst e) cse0 in let cse1' = Array.map (fun (ids, e) -> (ids, subst_expr subst e)) cse1 in diff --git a/tac2interp.ml b/tac2interp.ml index f93c8cb5fe..cd307e3ae7 100644 --- a/tac2interp.ml +++ b/tac2interp.ml @@ -72,9 +72,7 @@ let rec interp ist = function let iter (_, e) = e.clos_env <- ist in let () = List.iter iter fixs in interp ist e -| GTacTup [] -> - return (ValInt 0) -| GTacTup el | GTacArr el -> +| GTacArr el -> Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> return (ValBlk (0, Array.of_list el)) | GTacCst (_, n, []) -> return (ValInt n) diff --git a/tac2print.ml b/tac2print.ml index 96d0ceb875..a7f9ed48c8 100644 --- a/tac2print.ml +++ b/tac2print.ml @@ -165,7 +165,7 @@ let pr_glbexpr_gen lvl c = in let bnd = prlist_with_sep (fun () -> str "with" ++ spc ()) pr_bnd bnd in paren (str "let" ++ spc () ++ mut ++ bnd ++ str "in" ++ spc () ++ pr_glbexpr E5 e) - | GTacTup cl -> + | GTacCst (GCaseTuple _, _, cl) -> let paren = match lvl with | E0 | E1 -> paren | E2 | E3 | E4 | E5 -> fun x -> x @@ -173,7 +173,7 @@ let pr_glbexpr_gen lvl c = paren (prlist_with_sep (fun () -> str "," ++ spc ()) (pr_glbexpr E1) cl) | GTacArr cl -> mt () (** FIXME when implemented *) - | GTacCst (tpe, n, cl) -> + | GTacCst (GCaseAlg tpe, n, cl) -> begin match Tac2env.interp_type tpe with | _, GTydAlg def -> let paren = match lvl with -- cgit v1.2.3 From b5530d8953e74def1feb7dd651ba504e24749055 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 13 Dec 2016 16:23:13 +0100 Subject: Proper handling of exception definition in Ltac2. We actually implemented a full-fledged open type system, so that exceptions are a special case of it. --- Init.v | 2 +- g_ltac2.ml4 | 10 ++-- tac2core.ml | 10 +++- tac2entries.ml | 150 ++++++++++++++++++++++++++++++++++++++++++++++++++++- tac2entries.mli | 3 +- tac2env.ml | 5 +- tac2env.mli | 5 +- tac2expr.mli | 23 +++++++-- tac2intern.ml | 158 +++++++++++++++++++++++++++++++++++++++++++++----------- tac2intern.mli | 5 ++ tac2interp.ml | 38 ++++++++++++-- tac2interp.mli | 7 +-- tac2print.ml | 40 ++++++++++++-- 13 files changed, 398 insertions(+), 58 deletions(-) diff --git a/Init.v b/Init.v index 0f0b4636d8..322f275346 100644 --- a/Init.v +++ b/Init.v @@ -20,7 +20,7 @@ Ltac2 Type evar. Ltac2 Type constr. Ltac2 Type ident. Ltac2 Type message. -Ltac2 Type exn. +Ltac2 Type exn := [ .. ]. Ltac2 Type 'a array. (** Pervasive types *) diff --git a/g_ltac2.ml4 b/g_ltac2.ml4 index 9384584f19..a149d942c6 100644 --- a/g_ltac2.ml4 +++ b/g_ltac2.ml4 @@ -137,6 +137,7 @@ GEXTEND Gram ; tac2typ_knd: [ [ t = tac2type -> CTydDef (Some t) + | "["; ".."; "]" -> CTydOpn | "["; t = tac2alg_constructors; "]" -> CTydAlg t | "{"; t = tac2rec_fields; "}"-> CTydRec t ] ] ; @@ -173,9 +174,12 @@ GEXTEND Gram ] ] ; tac2typ_def: - [ [ prm = tac2typ_prm; id = locident; ":="; e = tac2typ_knd -> - (id, (prm, e)) - | prm = tac2typ_prm; id = locident -> (id, (prm, CTydDef None)) + [ [ prm = tac2typ_prm; id = Prim.qualid; (r, e) = tac2type_body -> (id, r, (prm, e)) ] ] + ; + tac2type_body: + [ [ -> false, CTydDef None + | ":="; e = tac2typ_knd -> false, e + | "::="; e = tac2typ_knd -> true, e ] ] ; tac2def_typ: diff --git a/tac2core.ml b/tac2core.ml index c18b6032a6..227bea8ddd 100644 --- a/tac2core.ml +++ b/tac2core.ml @@ -109,8 +109,14 @@ let to_constr c = to_ext val_constr c let of_ident c = of_ext val_ident c let to_ident c = to_ext val_ident c -let of_exn c = of_ext val_exn c -let to_exn c = to_ext val_exn c +(** FIXME: handle backtrace in Ltac2 exceptions *) +let of_exn c = match fst c with +| LtacError (kn, c) -> ValOpn (kn, c) +| _ -> of_ext val_exn c + +let to_exn c = match c with +| ValOpn (kn, c) -> (LtacError (kn, c), Exninfo.null) +| _ -> to_ext val_exn c let of_pp c = of_ext val_pp c let to_pp c = to_ext val_pp c diff --git a/tac2entries.ml b/tac2entries.ml index f4b3147c48..c776ad13d4 100644 --- a/tac2entries.ml +++ b/tac2entries.ml @@ -19,6 +19,8 @@ open Tac2print open Tac2intern open Vernacexpr +(** Tactic definition *) + type tacdef = { tacdef_local : bool; tacdef_expr : glb_tacexpr; @@ -52,6 +54,8 @@ let inTacDef : tacdef -> obj = subst_function = subst_tacdef; classify_function = classify_tacdef} +(** Type definition *) + type typdef = { typdef_local : bool; typdef_expr : glb_quant_typedef; @@ -86,6 +90,8 @@ let push_typedef visibility sp kn (_, def) = match def with in Tac2env.push_type visibility sp kn; List.iter iter fields +| GTydOpn -> + Tac2env.push_type visibility sp kn let next i = let ans = !i in @@ -108,7 +114,7 @@ let define_typedef kn (params, def as qdef) = match def with Tac2env.cdata_prms = params; cdata_type = kn; cdata_args = args; - cdata_indx = tag; + cdata_indx = Some tag; } in Tac2env.define_constructor knc data in @@ -129,6 +135,8 @@ let define_typedef kn (params, def as qdef) = match def with in Tac2env.define_type kn qdef; List.iteri iter fs +| GTydOpn -> + Tac2env.define_type kn qdef let perform_typdef vs ((sp, kn), def) = let () = if not def.typdef_local then push_typedef vs sp kn def.typdef_expr in @@ -155,6 +163,78 @@ let inTypDef : typdef -> obj = subst_function = subst_typdef; classify_function = classify_typdef} +(** Type extension *) + +type extension_data = { + edata_name : Id.t; + edata_args : int glb_typexpr list; +} + +type typext = { + typext_local : bool; + typext_prms : int; + typext_type : type_constant; + typext_expr : extension_data list; +} + +let push_typext vis sp kn def = + let iter data = + let spc = change_sp_label sp data.edata_name in + let knc = change_kn_label kn data.edata_name in + Tac2env.push_ltac vis spc (TacConstructor knc) + in + List.iter iter def.typext_expr + +let define_typext kn def = + let iter data = + let knc = change_kn_label kn data.edata_name in + let cdata = { + Tac2env.cdata_prms = def.typext_prms; + cdata_type = def.typext_type; + cdata_args = data.edata_args; + cdata_indx = None; + } in + Tac2env.define_constructor knc cdata + in + List.iter iter def.typext_expr + +let cache_typext ((sp, kn), def) = + let () = define_typext kn def in + push_typext (Until 1) sp kn def + +let perform_typext vs ((sp, kn), def) = + let () = if not def.typext_local then push_typext vs sp kn def in + define_typext kn def + +let load_typext i obj = perform_typext (Until i) obj +let open_typext i obj = perform_typext (Exactly i) obj + +let subst_typext (subst, e) = + let open Mod_subst in + let subst_data data = + let edata_args = List.smartmap (fun e -> subst_type subst e) data.edata_args in + if edata_args == data.edata_args then data + else { data with edata_args } + in + let typext_type = subst_kn subst e.typext_type in + let typext_expr = List.smartmap subst_data e.typext_expr in + if typext_type == e.typext_type && typext_expr == e.typext_expr then + e + else + { e with typext_type; typext_expr } + +let classify_typext o = Substitute o + +let inTypExt : typext -> obj = + declare_object {(default_object "TAC2-TYPE-EXTENSION") with + cache_function = cache_typext; + load_function = load_typext; + open_function = open_typext; + subst_function = subst_typext; + classify_function = classify_typext} + +(** Toplevel entries *) + let register_ltac ?(local = false) isrec tactics = if isrec then let map (na, e) = (na, None, e) in @@ -206,7 +286,12 @@ let register_ltac ?(local = false) isrec tactics = in List.iter iter defs -let register_type ?(local = false) isrec types = +let qualid_to_ident (loc, qid) = + let (dp, id) = Libnames.repr_qualid qid in + if DirPath.is_empty dp then (loc, id) + else user_err ~loc (str "Identifier expected") + +let register_typedef ?(local = false) isrec types = let same_name ((_, id1), _) ((_, id2), _) = Id.equal id1 id2 in let () = match List.duplicates same_name types with | [] -> () @@ -242,6 +327,10 @@ let register_type ?(local = false) isrec types = user_err (str "Multiple definitions of the projection " ++ Id.print id) in () + | CTydOpn -> + if isrec then + user_err ~loc (str "The open type declaration " ++ Id.print id ++ + str " cannot be recursive") in let () = List.iter check types in let self = @@ -289,6 +378,63 @@ let register_primitive ?(local = false) (loc, id) t ml = } in ignore (Lib.add_leaf id (inTacDef def)) +let register_open ?(local = false) (loc, qid) (params, def) = + let kn = + try Tac2env.locate_type qid + with Not_found -> + user_err ~loc (str "Unbound type " ++ pr_qualid qid) + in + let (tparams, t) = Tac2env.interp_type kn in + let () = match t with + | GTydOpn -> () + | GTydAlg _ | GTydRec _ | GTydDef _ -> + user_err ~loc (str "Type " ++ pr_qualid qid ++ str " is not an open type") + in + let () = + if not (Int.equal (List.length params) tparams) then + Tac2intern.error_nparams_mismatch loc (List.length params) tparams + in + match def with + | CTydOpn -> () + | CTydAlg def -> + let intern_type t = + let tpe = CTydDef (Some t) in + let (_, ans) = intern_typedef Id.Map.empty (params, tpe) in + match ans with + | GTydDef (Some t) -> t + | _ -> assert false + in + let map (id, tpe) = + let tpe = List.map intern_type tpe in + { edata_name = id; edata_args = tpe } + in + let def = List.map map def in + let def = { + typext_local = local; + typext_type = kn; + typext_prms = tparams; + typext_expr = def; + } in + Lib.add_anonymous_leaf (inTypExt def) + | CTydRec _ | CTydDef _ -> + user_err ~loc (str "Extensions only accept inductive constructors") + +let register_type ?local isrec types = match types with +| [qid, true, def] -> + let (loc, _) = qid in + let () = if isrec then user_err ~loc (str "Extensions cannot be recursive") in + register_open ?local qid def +| _ -> + let map (qid, redef, def) = + let (loc, _) = qid in + let () = if redef then + user_err ~loc (str "Types can only be extended one by one") + in + (qualid_to_ident qid, def) + in + let types = List.map map types in + register_typedef ?local isrec types + let register_struct ?local str = match str with | StrVal (isrec, e) -> register_ltac ?local isrec e | StrTyp (isrec, t) -> register_type ?local isrec t diff --git a/tac2entries.mli b/tac2entries.mli index 9c5d0a15fd..0d9b3ad134 100644 --- a/tac2entries.mli +++ b/tac2entries.mli @@ -8,6 +8,7 @@ open Loc open Names +open Libnames open Tac2expr (** {5 Toplevel definitions} *) @@ -16,7 +17,7 @@ val register_ltac : ?local:bool -> rec_flag -> (Name.t located * raw_tacexpr) list -> unit val register_type : ?local:bool -> rec_flag -> - (Id.t located * raw_quant_typedef) list -> unit + (qualid located * redef_flag * raw_quant_typedef) list -> unit val register_primitive : ?local:bool -> Id.t located -> raw_typexpr -> ml_tactic_name -> unit diff --git a/tac2env.ml b/tac2env.ml index 5e379473c8..0fcdba1ca7 100644 --- a/tac2env.ml +++ b/tac2env.ml @@ -22,7 +22,7 @@ type constructor_data = { cdata_prms : int; cdata_type : type_constant; cdata_args : int glb_typexpr list; - cdata_indx : int; + cdata_indx : int option; } type projection_data = { @@ -62,8 +62,9 @@ let rec eval_pure = function ValCls { clos_env = Id.Map.empty; clos_var = na; clos_exp = e } | GTacCst (_, n, []) -> ValInt n | GTacCst (_, n, el) -> ValBlk (n, Array.map_of_list eval_pure el) +| GTacOpn (kn, el) -> ValOpn (kn, Array.map_of_list eval_pure el) | GTacAtm (AtmStr _) | GTacArr _ | GTacLet _ | GTacVar _ | GTacSet _ -| GTacApp _ | GTacCse _ | GTacPrj _ | GTacPrm _ | GTacExt _ -> +| GTacApp _ | GTacCse _ | GTacPrj _ | GTacPrm _ | GTacExt _ | GTacWth _ -> anomaly (Pp.str "Term is not a syntactical value") let define_global kn e = diff --git a/tac2env.mli b/tac2env.mli index 16232ec810..477f4ebec7 100644 --- a/tac2env.mli +++ b/tac2env.mli @@ -33,10 +33,11 @@ type constructor_data = { (** Inductive definition to which the constructor pertains *) cdata_args : int glb_typexpr list; (** Types of the constructor arguments *) - cdata_indx : int; + cdata_indx : int option; (** Index of the constructor in the ADT. Numbering is duplicated between argumentless and argument-using constructors, e.g. in type ['a option] - [None] and [Some] have both index 0. *) + [None] and [Some] have both index 0. This field is empty whenever the + constructor is a member of an open type. *) } val define_constructor : ltac_constructor -> constructor_data -> unit diff --git a/tac2expr.mli b/tac2expr.mli index 1fac5a2315..1840b567b4 100644 --- a/tac2expr.mli +++ b/tac2expr.mli @@ -13,6 +13,7 @@ open Libnames type mutable_flag = bool type rec_flag = bool +type redef_flag = bool type lid = Id.t type uid = Id.t @@ -40,6 +41,7 @@ type raw_typedef = | CTydDef of raw_typexpr option | CTydAlg of (uid * raw_typexpr list) list | CTydRec of (lid * mutable_flag * raw_typexpr) list +| CTydOpn type 'a glb_typexpr = | GTypVar of 'a @@ -51,6 +53,7 @@ type glb_typedef = | GTydDef of int glb_typexpr option | GTydAlg of (uid * int glb_typexpr list) list | GTydRec of (lid * mutable_flag * int glb_typexpr) list +| GTydOpn type type_scheme = int * int glb_typexpr @@ -94,6 +97,13 @@ type case_info = | GCaseTuple of int | GCaseAlg of type_constant +type 'a open_match = { + opn_match : 'a; + opn_branch : (Name.t * Name.t array * 'a) KNmap.t; + (** Invariant: should not be empty *) + opn_default : Name.t * 'a; +} + type glb_tacexpr = | GTacAtm of atom | GTacVar of Id.t @@ -106,22 +116,27 @@ type glb_tacexpr = | GTacCse of glb_tacexpr * case_info * glb_tacexpr array * (Name.t array * glb_tacexpr) array | GTacPrj of type_constant * glb_tacexpr * int | GTacSet of type_constant * glb_tacexpr * int * glb_tacexpr +| GTacOpn of ltac_constructor * glb_tacexpr list +| GTacWth of glb_tacexpr open_match | GTacExt of glob_generic_argument | GTacPrm of ml_tactic_name * glb_tacexpr list (** Toplevel statements *) type strexpr = | StrVal of rec_flag * (Name.t located * raw_tacexpr) list -| StrTyp of rec_flag * (Id.t located * raw_quant_typedef) list + (** Term definition *) +| StrTyp of rec_flag * (qualid located * redef_flag * raw_quant_typedef) list + (** Type definition *) | StrPrm of Id.t located * raw_typexpr * ml_tactic_name + (** External definition *) (** {5 Dynamic semantics} *) (** Values are represented in a way similar to OCaml, i.e. they constrast immediate integers (integers, constructors without arguments) and structured blocks (tuples, arrays, constructors with arguments), as well as a few other - base cases, namely closures, strings and dynamic type coming from the Coq - implementation. *) + base cases, namely closures, strings, named constructors, and dynamic type + coming from the Coq implementation. *) type tag = int @@ -134,6 +149,8 @@ type valexpr = (** Strings *) | ValCls of closure (** Closures *) +| ValOpn of KerName.t * valexpr array + (** Open constructors *) | ValExt of Geninterp.Val.t (** Arbitrary data *) diff --git a/tac2intern.ml b/tac2intern.ml index c4d2fc277b..32d1e07c17 100644 --- a/tac2intern.ml +++ b/tac2intern.ml @@ -215,6 +215,11 @@ let error_nargs_mismatch loc nargs nfound = str " arguments, but is applied to " ++ int nfound ++ str " arguments") +let error_nparams_mismatch loc nargs nfound = + user_err ~loc (str "Type expects " ++ int nargs ++ + str " arguments, but is applied to " ++ int nfound ++ + str " arguments") + let rec subst_type subst (t : 'a glb_typexpr) = match t with | GTypVar id -> subst id | GTypArrow (t1, t2) -> GTypArrow (subst_type subst t1, subst_type subst t2) @@ -275,7 +280,7 @@ let fresh_reftype env (kn : KerName.t) = let is_unfoldable kn = match snd (Tac2env.interp_type kn) with | GTydDef (Some _) -> true -| GTydDef None | GTydAlg _ | GTydRec _ -> false +| GTydDef None | GTydAlg _ | GTydRec _ | GTydOpn -> false let unfold env kn args = let (nparams, def) = Tac2env.interp_type kn in @@ -349,7 +354,7 @@ let unify loc env t1 t2 = let is_pure_constructor kn = match snd (Tac2env.interp_type kn) with - | GTydAlg _ -> true + | GTydAlg _ | GTydOpn -> true | GTydRec fields -> let is_pure (_, mut, _) = not mut in List.for_all is_pure fields @@ -360,14 +365,16 @@ let rec is_value = function | GTacAtm (AtmStr _) | GTacApp _ | GTacLet _ -> false | GTacCst (GCaseTuple _, _, el) -> List.for_all is_value el | GTacCst (_, _, []) -> true +| GTacOpn (_, el) -> List.for_all is_value el | GTacCst (GCaseAlg kn, _, el) -> is_pure_constructor kn && List.for_all is_value el -| GTacArr _ | GTacCse _ | GTacPrj _ | GTacSet _ | GTacExt _ | GTacPrm _ -> false +| GTacArr _ | GTacCse _ | GTacPrj _ | GTacSet _ | GTacExt _ | GTacPrm _ +| GTacWth _ -> false let is_rec_rhs = function | GTacFun _ -> true | GTacAtm _ | GTacVar _ | GTacRef _ | GTacApp _ | GTacLet _ | GTacPrj _ | GTacSet _ | GTacArr _ | GTacExt _ | GTacPrm _ | GTacCst _ -| GTacCse _ -> false +| GTacCse _ | GTacOpn _ | GTacWth _ -> false let rec fv_type f t accu = match t with | GTypVar id -> f id accu @@ -503,40 +510,55 @@ let invalid_pattern ~loc kn t = user_err ~loc (str "Invalid pattern, expected a pattern for type " ++ pr_typref kn ++ str ", found a pattern of type " ++ pt) (** FIXME *) +(** Pattern view *) + +type glb_patexpr = +| GPatVar of Name.t +| GPatRef of ltac_constructor * glb_patexpr list +| GPatTup of glb_patexpr list + +let rec intern_patexpr env = function +| CPatAny _ -> GPatVar Anonymous +| CPatRef (_, qid, []) -> + begin match get_constructor env qid with + | ArgVar (_, id) -> GPatVar (Name id) + | ArgArg (_, kn) -> GPatRef (kn, []) + end +| CPatRef (_, qid, pl) -> + begin match get_constructor env qid with + | ArgVar (loc, _) -> + user_err ~loc (str "Unbound constructor " ++ pr_qualid (snd qid)) + | ArgArg (_, kn) -> GPatRef (kn, List.map (fun p -> intern_patexpr env p) pl) + end +| CPatTup (_, pl) -> + GPatTup (List.map (fun p -> intern_patexpr env p) pl) + type pattern_kind = | PKind_empty -| PKind_variant of KerName.t +| PKind_variant of type_constant +| PKind_open of type_constant | PKind_tuple of int | PKind_any let get_pattern_kind env pl = match pl with | [] -> PKind_empty | p :: pl -> - let rec get_kind p pl = match fst p with - | CPatAny _ -> + let rec get_kind (p, _) pl = match intern_patexpr env p with + | GPatVar _ -> begin match pl with | [] -> PKind_any | p :: pl -> get_kind p pl end - | CPatRef (_, qid, []) -> - begin match get_constructor env qid with - | ArgVar _ -> - begin match pl with - | [] -> PKind_any - | p :: pl -> get_kind p pl - end - | ArgArg (data, _) -> PKind_variant data.cdata_type - end - | CPatRef (_, qid, _ :: _) -> - begin match get_constructor env qid with - | ArgVar (loc, _) -> - user_err ~loc (str "Unbound constructor " ++ pr_qualid (snd qid)) - | ArgArg (data, _) -> PKind_variant data.cdata_type - end - | CPatTup (_, tp) -> PKind_tuple (List.length tp) + | GPatRef (kn, pl) -> + let data = Tac2env.interp_constructor kn in + if Option.is_empty data.cdata_indx then PKind_open data.cdata_type + else PKind_variant data.cdata_type + | GPatTup tp -> PKind_tuple (List.length tp) in get_kind p pl +(** Internalization *) + let is_constructor env qid = match get_variable env qid with | ArgArg (TacConstructor _) -> true | _ -> false @@ -739,9 +761,8 @@ and intern_case env loc e pl = match get_pattern_kind env pl with | PKind_any -> let (pat, b) = List.hd pl in - let na = match pat with - | CPatAny _ -> Anonymous - | CPatRef (_, (_, qid), _) -> Name (snd (repr_qualid qid)) + let na = match intern_patexpr env pat with + | GPatVar na -> na | _ -> assert false in let () = check_redundant_clause (List.tl pl) in @@ -851,7 +872,10 @@ and intern_case env loc e pl = let nenv = List.fold_left2 fold env ids data.cdata_args in let (br', brT) = intern_rec nenv br in let () = - let index = data.cdata_indx in + let index = match data.cdata_indx with + | Some i -> i + | None -> assert false + in if List.is_empty args then if Option.is_empty const.(index) then const.(index) <- Some br' else warn_redundant_clause ~loc () @@ -869,13 +893,67 @@ and intern_case env loc e pl = in let () = intern_branch pl in let map = function - | None -> user_err ~loc (str "Unhandled match case") (** FIXME *) + | None -> user_err ~loc (str "TODO: Unhandled match case") (** FIXME *) | Some x -> x in let const = Array.map map const in let nonconst = Array.map map nonconst in let ce = GTacCse (e', GCaseAlg kn, const, nonconst) in (ce, ret) + | PKind_open kn -> + let subst, tc = fresh_reftype env kn in + let () = unify (loc_of_tacexpr e) env t tc in + let ret = GTypVar (fresh_id env) in + let rec intern_branch map = function + | [] -> + user_err ~loc (str "Missing default case") + | (pat, br) :: rem -> + match intern_patexpr env pat with + | GPatVar na -> + let () = check_redundant_clause rem in + let nenv = push_name na (monomorphic tc) env in + let br' = intern_rec_with_constraint nenv br ret in + let def = (na, br') in + (map, def) + | GPatRef (knc, args) -> + let get = function + | GPatVar na -> na + | GPatRef _ | GPatTup _ -> + user_err ~loc (str "TODO: Unhandled match case") (** FIXME *) + in + let loc = loc_of_patexpr pat in + let ids = List.map get args in + let data = Tac2env.interp_constructor knc in + let () = + if not (KerName.equal kn data.cdata_type) then + invalid_pattern ~loc kn (GCaseAlg data.cdata_type) + in + let nids = List.length ids in + let nargs = List.length data.cdata_args in + let () = + if not (Int.equal nids nargs) then error_nargs_mismatch loc nargs nids + in + let fold env id tpe = + (** Instantiate all arguments *) + let subst n = GTypVar subst.(n) in + let tpe = subst_type subst tpe in + push_name id (monomorphic tpe) env + in + let nenv = List.fold_left2 fold env ids data.cdata_args in + let br' = intern_rec_with_constraint nenv br ret in + let map = + if KNmap.mem knc map then + let () = warn_redundant_clause ~loc () in + map + else + KNmap.add knc (Anonymous, Array.of_list ids, br') map + in + intern_branch map rem + | GPatTup tup -> + invalid_pattern ~loc kn (GCaseTuple (List.length tup)) + in + let (map, def) = intern_branch KNmap.empty pl in + (GTacWth { opn_match = e'; opn_branch = map; opn_default = def }, ret) and intern_constructor env loc kn args = let cstr = interp_constructor kn in @@ -887,7 +965,11 @@ and intern_constructor env loc kn args = let ans = GTypRef (cstr.cdata_type, List.init cstr.cdata_prms (fun i -> GTypVar subst.(i))) in let map arg tpe = intern_rec_with_constraint env arg tpe in let args = List.map2 map args types in - (GTacCst (GCaseAlg cstr.cdata_type, cstr.cdata_indx, args), ans) + match cstr.cdata_indx with + | Some idx -> + (GTacCst (GCaseAlg cstr.cdata_type, idx, args), ans) + | None -> + (GTacOpn (kn, args), ans) else error_nargs_mismatch loc nargs (List.length args) @@ -986,6 +1068,7 @@ let intern_typedef self (ids, t) : glb_quant_typedef = let map (c, mut, t) = (c, mut, intern t) in let fields = List.map map fields in (count, GTydRec fields) + | CTydOpn -> (count, GTydOpn) let intern_open_type t = let env = empty_env () in @@ -1045,6 +1128,18 @@ let rec subst_expr subst e = match e with let cse1' = Array.map (fun (ids, e) -> (ids, subst_expr subst e)) cse1 in let ci' = subst_case_info subst ci in GTacCse (subst_expr subst e, ci', cse0', cse1') +| GTacWth { opn_match = e; opn_branch = br; opn_default = (na, def) } as e0 -> + let e' = subst_expr subst e in + let def' = subst_expr subst def in + let fold kn (self, vars, p) accu = + let kn' = subst_kn subst kn in + let p' = subst_expr subst p in + if kn' == kn && p' == p then accu + else KNmap.add kn' (self, vars, p') (KNmap.remove kn accu) + in + let br' = KNmap.fold fold br br in + if e' == e && br' == br && def' == def then e0 + else GTacWth { opn_match = e'; opn_default = (na, def'); opn_branch = br' } | GTacPrj (kn, e, p) as e0 -> let kn' = subst_kn subst kn in let e' = subst_expr subst e in @@ -1057,6 +1152,10 @@ let rec subst_expr subst e = match e with | GTacExt ext -> let ext' = Genintern.generic_substitute subst ext in if ext' == ext then e else GTacExt ext' +| GTacOpn (kn, el) as e0 -> + let kn' = subst_kn subst kn in + let el' = List.smartmap (fun e -> subst_expr subst e) el in + if kn' == kn && el' == el then e0 else GTacOpn (kn', el') let subst_typedef subst e = match e with | GTydDef t -> @@ -1076,6 +1175,7 @@ let subst_typedef subst e = match e with in let fields' = List.smartmap map fields in if fields' == fields then e else GTydRec fields' +| GTydOpn -> GTydOpn let subst_quant_typedef subst (prm, def as qdef) = let def' = subst_typedef subst def in diff --git a/tac2intern.mli b/tac2intern.mli index 6633792e7e..4c482d0b0c 100644 --- a/tac2intern.mli +++ b/tac2intern.mli @@ -26,3 +26,8 @@ val subst_type : substitution -> 'a glb_typexpr -> 'a glb_typexpr val subst_expr : substitution -> glb_tacexpr -> glb_tacexpr val subst_quant_typedef : substitution -> glb_quant_typedef -> glb_quant_typedef val subst_type_scheme : substitution -> type_scheme -> type_scheme + +(** Errors *) + +val error_nargs_mismatch : Loc.t -> int -> int -> 'a +val error_nparams_mismatch : Loc.t -> int -> int -> 'a diff --git a/tac2interp.ml b/tac2interp.ml index cd307e3ae7..664b7de3d6 100644 --- a/tac2interp.ml +++ b/tac2interp.ml @@ -14,7 +14,14 @@ open Names open Proofview.Notations open Tac2expr -exception LtacError of KerName.t * valexpr +exception LtacError of KerName.t * valexpr array + +let () = register_handler begin function +| LtacError (kn, _) -> + let c = Tac2print.pr_constructor kn in + hov 0 (str "Uncaught Ltac2 exception:" ++ spc () ++ hov 0 c) +| _ -> raise Unhandled +end let val_exn = Geninterp.Val.create "ltac2:exn" @@ -81,12 +88,17 @@ let rec interp ist = function return (ValBlk (n, Array.of_list el)) | GTacCse (e, _, cse0, cse1) -> interp ist e >>= fun e -> interp_case ist e cse0 cse1 +| GTacWth { opn_match = e; opn_branch = cse; opn_default = def } -> + interp ist e >>= fun e -> interp_with ist e cse def | GTacPrj (_, e, p) -> interp ist e >>= fun e -> interp_proj ist e p | GTacSet (_, e, p, r) -> interp ist e >>= fun e -> interp ist r >>= fun r -> interp_set ist e p r +| GTacOpn (kn, el) -> + Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> + return (ValOpn (kn, Array.of_list el)) | GTacPrm (ml, el) -> Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> Tac2env.interp_primitive ml el @@ -106,7 +118,7 @@ and interp_app f args = match f with | id :: ids, arg :: args -> push (push_name ist id arg) ids args in push ist ids args -| ValExt _ | ValInt _ | ValBlk _ | ValStr _ -> +| ValExt _ | ValInt _ | ValBlk _ | ValStr _ | ValOpn _ -> anomaly (str "Unexpected value shape") and interp_case ist e cse0 cse1 = match e with @@ -115,18 +127,34 @@ and interp_case ist e cse0 cse1 = match e with let (ids, e) = cse1.(n) in let ist = CArray.fold_left2 push_name ist ids args in interp ist e -| ValExt _ | ValStr _ | ValCls _ -> +| ValExt _ | ValStr _ | ValCls _ | ValOpn _ -> + anomaly (str "Unexpected value shape") + +and interp_with ist e cse def = match e with +| ValOpn (kn, args) -> + let br = try Some (KNmap.find kn cse) with Not_found -> None in + begin match br with + | None -> + let (self, def) = def in + let ist = push_name ist self e in + interp ist def + | Some (self, ids, p) -> + let ist = push_name ist self e in + let ist = CArray.fold_left2 push_name ist ids args in + interp ist p + end +| ValInt _ | ValBlk _ | ValExt _ | ValStr _ | ValCls _ -> anomaly (str "Unexpected value shape") and interp_proj ist e p = match e with | ValBlk (_, args) -> return args.(p) -| ValInt _ | ValExt _ | ValStr _ | ValCls _ -> +| ValInt _ | ValExt _ | ValStr _ | ValCls _ | ValOpn _ -> anomaly (str "Unexpected value shape") and interp_set ist e p r = match e with | ValBlk (_, args) -> let () = args.(p) <- r in return (ValInt 0) -| ValInt _ | ValExt _ | ValStr _ | ValCls _ -> +| ValInt _ | ValExt _ | ValStr _ | ValCls _ | ValOpn _ -> anomaly (str "Unexpected value shape") diff --git a/tac2interp.mli b/tac2interp.mli index 7fe78a9460..bf6b2d4dde 100644 --- a/tac2interp.mli +++ b/tac2interp.mli @@ -20,8 +20,9 @@ val interp_app : valexpr -> valexpr list -> valexpr Proofview.tactic (** {5 Exceptions} *) -exception LtacError of KerName.t * valexpr -(** Ltac2-defined exceptions *) +exception LtacError of KerName.t * valexpr array +(** Ltac2-defined exceptions seen from OCaml side *) val val_exn : Exninfo.iexn Geninterp.Val.typ -(** Toplevel representation of Ltac2 exceptions *) +(** Toplevel representation of OCaml exceptions. Invariant: no [LtacError] + should be put into a value with tag [val_exn]. *) diff --git a/tac2print.ml b/tac2print.ml index a7f9ed48c8..ffa5ddc05a 100644 --- a/tac2print.ml +++ b/tac2print.ml @@ -153,7 +153,7 @@ let pr_glbexpr_gen lvl c = | E0 -> paren | E1 | E2 | E3 | E4 | E5 -> fun x -> x in - paren (pr_glbexpr E1 c) ++ spc () ++ (pr_sequence (pr_glbexpr E0) cl) + paren (pr_glbexpr E1 c ++ spc () ++ (pr_sequence (pr_glbexpr E0) cl)) | GTacLet (mut, bnd, e) -> let paren = match lvl with | E0 | E1 | E2 | E3 | E4 -> paren @@ -195,7 +195,7 @@ let pr_glbexpr_gen lvl c = in let args = prlist_with_sep (fun () -> str ";" ++ spc ()) pr_arg args in str "{" ++ spc () ++ args ++ spc () ++ str "}" - | _, GTydDef _ -> assert false + | _, (GTydDef _ | GTydOpn) -> assert false end | GTacCse (e, info, cst_br, ncst_br) -> let e = pr_glbexpr E5 e in @@ -203,7 +203,7 @@ let pr_glbexpr_gen lvl c = | GCaseAlg kn -> let def = match Tac2env.interp_type kn with | _, GTydAlg def -> def - | _, GTydDef _ | _, GTydRec _ -> assert false + | _, GTydDef _ | _, GTydRec _ | _, GTydOpn -> assert false in let br = order_branches cst_br ncst_br def in let pr_branch (cstr, vars, p) = @@ -224,10 +224,33 @@ let pr_glbexpr_gen lvl c = str "|" ++ spc () ++ paren vars ++ spc () ++ str "=>" ++ spc () ++ p in hov 0 (hov 0 (str "match" ++ spc () ++ e ++ spc () ++ str "with") ++ spc () ++ Pp.v 0 br ++ str "end") + | GTacWth wth -> + let e = pr_glbexpr E5 wth.opn_match in + let pr_pattern c self vars p = + let self = match self with + | Anonymous -> mt () + | Name id -> spc () ++ str "as" ++ spc () ++ Id.print id + in + hov 0 (str "|" ++ spc () ++ c ++ vars ++ self ++ spc () ++ str "=>" ++ spc () ++ + hov 2 (pr_glbexpr E5 p)) ++ spc () + in + let pr_branch (cstr, (self, vars, p)) = + let cstr = pr_constructor cstr in + let vars = match Array.to_list vars with + | [] -> mt () + | vars -> spc () ++ pr_sequence pr_name vars + in + pr_pattern cstr self vars p + in + let br = prlist pr_branch (KNmap.bindings wth.opn_branch) in + let (def_as, def_p) = wth.opn_default in + let def = pr_pattern (str "_") def_as (mt ()) def_p in + let br = br ++ def in + hov 0 (hov 0 (str "match" ++ spc () ++ e ++ spc () ++ str "with") ++ spc () ++ Pp.v 0 br ++ str "end") | GTacPrj (kn, e, n) -> let def = match Tac2env.interp_type kn with | _, GTydRec def -> def - | _, GTydDef _ | _, GTydAlg _ -> assert false + | _, GTydDef _ | _, GTydAlg _ | _, GTydOpn -> assert false in let (proj, _, _) = List.nth def n in let proj = change_kn_label kn proj in @@ -237,7 +260,7 @@ let pr_glbexpr_gen lvl c = | GTacSet (kn, e, n, r) -> let def = match Tac2env.interp_type kn with | _, GTydRec def -> def - | _, GTydDef _ | _, GTydAlg _ -> assert false + | _, GTydDef _ | _, GTydAlg _ | _, GTydOpn -> assert false in let (proj, _, _) = List.nth def n in let proj = change_kn_label kn proj in @@ -245,6 +268,13 @@ let pr_glbexpr_gen lvl c = let e = pr_glbexpr E0 e in let r = pr_glbexpr E1 r in e ++ str "." ++ paren proj ++ spc () ++ str ":=" ++ spc () ++ r + | GTacOpn (kn, cl) -> + let paren = match lvl with + | E0 -> paren + | E1 | E2 | E3 | E4 | E5 -> fun x -> x + in + let c = pr_constructor kn in + paren (c ++ spc () ++ (pr_sequence (pr_glbexpr E0) cl)) | GTacExt arg -> let GenArg (Glbwit tag, arg) = arg in let name = match tag with -- cgit v1.2.3 From 735ab0a7d2f7afaed0695e014034f4b2d6e287c8 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 14 Dec 2016 08:09:54 +0100 Subject: Stdlib functions now return Ltac2 exceptions. --- Init.v | 14 ++++++++++++++ Int.v | 2 ++ tac2core.ml | 55 ++++++++++++++++++++++++++++++++----------------------- 3 files changed, 48 insertions(+), 23 deletions(-) diff --git a/Init.v b/Init.v index 322f275346..8ff5837bb4 100644 --- a/Init.v +++ b/Init.v @@ -30,3 +30,17 @@ Ltac2 Type 'a option := [ None | Some ('a) ]. Ltac2 Type 'a ref := { mutable contents : 'a }. Ltac2 Type bool := [ true | false ]. + +(** Pervasive exceptions *) + +Ltac2 Type exn ::= [ Out_of_bounds ]. +(** Used for bound checking, e.g. with String and Array. *) + +Ltac2 Type exn ::= [ Not_focussed ]. +(** In Ltac2, the notion of "current environment" only makes sense when there is + at most one goal under focus. Contrarily to Ltac1, instead of dynamically + focussing when we need it, we raise this non-backtracking error when it does + not make sense. *) + +Ltac2 Type exn ::= [ Not_found ]. +(** Used when something is missing. *) diff --git a/Int.v b/Int.v index bb0287efc8..ab43eb27b0 100644 --- a/Int.v +++ b/Int.v @@ -8,6 +8,8 @@ Require Import Coq.ltac2.Init. +Ltac2 Type exn ::= [ Division_by_zero ]. + Ltac2 @ external equal : int -> int -> bool := "ltac2" "int_equal". Ltac2 @ external compare : int -> int -> int := "ltac2" "int_compare". Ltac2 @ external add : int -> int -> int := "ltac2" "int_add". diff --git a/tac2core.ml b/tac2core.ml index 227bea8ddd..3232fcba5b 100644 --- a/tac2core.ml +++ b/tac2core.ml @@ -20,6 +20,10 @@ open Proofview.Notations (** Standard values *) let coq_core n = KerName.make2 Tac2env.coq_prefix (Label.of_id (Id.of_string_soft n)) +let stdlib_prefix md = + MPfile (DirPath.make (List.map Id.of_string [md; "ltac2"; "Coq"])) +let coq_stdlib md n = + KerName.make2 (stdlib_prefix md) (Label.of_id (Id.of_string n)) let val_tag t = match val_tag t with | Val.Base t -> t @@ -125,9 +129,21 @@ end let val_valexpr = Val.create "ltac2:valexpr" +(** Stdlib exceptions *) + +let err_notfocussed = + LtacError (coq_core "Not_focussed", [||]) + +let err_outofbounds = + LtacError (coq_core "Out_of_bounds", [||]) + +let err_notfound = + LtacError (coq_core "Not_found", [||]) + (** Helper functions *) let thaw f = interp_app f [v_unit] +let throw e = Proofview.tclLIFT (Proofview.NonLogical.raise e) let return x = Proofview.tclUNIT x let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } @@ -138,16 +154,9 @@ let wrap f = let wrap_unit f = return () >>= fun () -> f (); return v_unit -(** In Ltac2, the notion of "current environment" only makes sense when there is - at most one goal under focus. Contrarily to Ltac1, instead of dynamically - focussing when we need it, we raise a non-backtracking error when it does - not make sense. *) -exception NonFocussedGoal - -let () = register_handler begin function -| NonFocussedGoal -> str "Several goals under focus" -| _ -> raise Unhandled -end +let wrap_exn f err = + return () >>= fun () -> + try return (f ()) with e when CErrors.noncritical e -> err let pf_apply f = Proofview.Goal.goals >>= function @@ -159,7 +168,7 @@ let pf_apply f = gl >>= fun gl -> f (Proofview.Goal.env gl) (Tacmach.New.project gl) | _ :: _ :: _ -> - Proofview.tclLIFT (Proofview.NonLogical.raise NonFocussedGoal) + throw err_notfocussed (** Primitives *) @@ -197,8 +206,8 @@ let prm_message_concat : ml_tactic = function let prm_array_make : ml_tactic = function | [ValInt n; x] -> - (** FIXME: wrap exception *) - wrap (fun () -> ValBlk (0, Array.make n x)) + if n < 0 || n > Sys.max_array_length then throw err_outofbounds + else wrap (fun () -> ValBlk (0, Array.make n x)) | _ -> assert false let prm_array_length : ml_tactic = function @@ -207,14 +216,14 @@ let prm_array_length : ml_tactic = function let prm_array_set : ml_tactic = function | [ValBlk (_, v); ValInt n; x] -> - (** FIXME: wrap exception *) - wrap_unit (fun () -> v.(n) <- x) + if n < 0 || n >= Array.length v then throw err_outofbounds + else wrap_unit (fun () -> v.(n) <- x) | _ -> assert false let prm_array_get : ml_tactic = function | [ValBlk (_, v); ValInt n] -> - (** FIXME: wrap exception *) - wrap (fun () -> v.(n)) + if n < 0 || n >= Array.length v then throw err_outofbounds + else wrap (fun () -> v.(n)) | _ -> assert false (** Int *) @@ -243,8 +252,8 @@ let prm_string_make : ml_tactic = function | [n; c] -> let n = Value.to_int n in let c = Value.to_char c in - (** FIXME: wrap exception *) - wrap (fun () -> Value.of_string (Bytes.make n c)) + if n < 0 || n > Sys.max_string_length then throw err_outofbounds + else wrap (fun () -> Value.of_string (Bytes.make n c)) | _ -> assert false let prm_string_length : ml_tactic = function @@ -257,16 +266,16 @@ let prm_string_set : ml_tactic = function let s = Value.to_string s in let n = Value.to_int n in let c = Value.to_char c in - (** FIXME: wrap exception *) - wrap_unit (fun () -> Bytes.set s n c) + if n < 0 || n >= Bytes.length s then throw err_outofbounds + else wrap_unit (fun () -> Bytes.set s n c) | _ -> assert false let prm_string_get : ml_tactic = function | [s; n] -> let s = Value.to_string s in let n = Value.to_int n in - (** FIXME: wrap exception *) - wrap (fun () -> Value.of_char (Bytes.get s n)) + if n < 0 || n >= Bytes.length s then throw err_outofbounds + else wrap (fun () -> Value.of_char (Bytes.get s n)) | _ -> assert false (** Error *) -- cgit v1.2.3 From c341a00d916c27b75c79c2fdcce13e969772a990 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 16 Mar 2017 09:05:04 +0100 Subject: Allow raw terms to contain references to absolute definitions. --- g_ltac2.ml4 | 20 +++++++++---------- tac2entries.ml | 2 +- tac2env.ml | 7 +------ tac2env.mli | 4 ---- tac2expr.mli | 20 +++++++++++++------ tac2intern.ml | 62 ++++++++++++++++++++++++++++++++++++++++++---------------- 6 files changed, 71 insertions(+), 44 deletions(-) diff --git a/g_ltac2.ml4 b/g_ltac2.ml4 index a149d942c6..30f5399a88 100644 --- a/g_ltac2.ml4 +++ b/g_ltac2.ml4 @@ -32,12 +32,12 @@ GEXTEND Gram GLOBAL: tac2expr tac2type tac2def_val tac2def_typ tac2def_ext; tac2pat: [ "1" LEFTA - [ id = Prim.qualid; pl = LIST1 tac2pat LEVEL "0" -> CPatRef (!@loc, id, pl) - | id = Prim.qualid -> CPatRef (!@loc, id, []) ] + [ id = Prim.qualid; pl = LIST1 tac2pat LEVEL "0" -> CPatRef (!@loc, RelId id, pl) + | id = Prim.qualid -> CPatRef (!@loc, RelId id, []) ] | "0" [ "_" -> CPatAny (!@loc) | "()" -> CPatTup (!@loc, []) - | id = Prim.qualid -> CPatRef (!@loc, id, []) + | id = Prim.qualid -> CPatRef (!@loc, RelId id, []) | "("; pl = LIST0 tac2pat LEVEL "1" SEP ","; ")" -> CPatTup (!@loc, pl) ] ] ; @@ -54,8 +54,8 @@ GEXTEND Gram [ e1 = tac2expr; ";"; e2 = tac2expr -> CTacSeq (!@loc, e1, e2) ] | "1" LEFTA [ e = tac2expr; el = LIST1 tac2expr LEVEL "0" -> CTacApp (!@loc, e, el) - | e = SELF; ".("; qid = Prim.qualid; ")" -> CTacPrj (!@loc, e, qid) - | e = SELF; ".("; qid = Prim.qualid; ")"; ":="; r = tac2expr LEVEL "1" -> CTacSet (!@loc, e, qid, r) + | e = SELF; ".("; qid = Prim.qualid; ")" -> CTacPrj (!@loc, e, RelId qid) + | e = SELF; ".("; qid = Prim.qualid; ")"; ":="; r = tac2expr LEVEL "1" -> CTacSet (!@loc, e, RelId qid, r) | e0 = tac2expr; ","; el = LIST1 tac2expr LEVEL "1" SEP "," -> CTacTup (!@loc, e0 :: el) ] | "0" [ "("; a = tac2expr LEVEL "5"; ")" -> a @@ -86,7 +86,7 @@ GEXTEND Gram tactic_atom: [ [ n = Prim.integer -> CTacAtm (!@loc, AtmInt n) | s = Prim.string -> CTacAtm (!@loc, AtmStr s) - | id = Prim.qualid -> CTacRef id + | id = Prim.qualid -> CTacRef (RelId id) | IDENT "constr"; ":"; "("; c = Constr.lconstr; ")" -> inj_constr !@loc c | IDENT "open_constr"; ":"; "("; c = Constr.lconstr; ")" -> inj_open_constr !@loc c | IDENT "ident"; ":"; "("; c = Prim.ident; ")" -> inj_ident !@loc c @@ -104,14 +104,14 @@ GEXTEND Gram | "2" [ t = tac2type; "*"; tl = LIST1 tac2type SEP "*" -> CTypTuple (!@loc, t :: tl) ] | "1" LEFTA - [ t = SELF; qid = Prim.qualid -> CTypRef (!@loc, qid, [t]) ] + [ t = SELF; qid = Prim.qualid -> CTypRef (!@loc, RelId qid, [t]) ] | "0" [ "("; t = tac2type LEVEL "5"; ")" -> t | id = typ_param -> CTypVar (!@loc, Name id) | "_" -> CTypVar (!@loc, Anonymous) - | qid = Prim.qualid -> CTypRef (!@loc, qid, []) + | qid = Prim.qualid -> CTypRef (!@loc, RelId qid, []) | "("; p = LIST1 tac2type LEVEL "5" SEP ","; ")"; qid = Prim.qualid -> - CTypRef (!@loc, qid, p) ] + CTypRef (!@loc, RelId qid, p) ] ]; locident: [ [ id = Prim.ident -> (!@loc, id) ] ] @@ -165,7 +165,7 @@ GEXTEND Gram | -> [] ] ] ; tac2rec_fieldexpr: - [ [ qid = Prim.qualid; ":="; e = tac2expr LEVEL "1" -> qid, e ] ] + [ [ qid = Prim.qualid; ":="; e = tac2expr LEVEL "1" -> RelId qid, e ] ] ; tac2typ_prm: [ [ -> [] diff --git a/tac2entries.ml b/tac2entries.ml index c776ad13d4..374a367188 100644 --- a/tac2entries.ml +++ b/tac2entries.ml @@ -243,7 +243,7 @@ let register_ltac ?(local = false) isrec tactics = | Anonymous -> None | Name id -> let qid = Libnames.qualid_of_ident id in - let e = CTacLet (Loc.ghost, true, bindings, CTacRef (loc, qid)) in + let e = CTacLet (Loc.ghost, true, bindings, CTacRef (RelId (loc, qid))) in let (e, t) = intern e in let e = match e with | GTacLet (true, _, e) -> assert false diff --git a/tac2env.ml b/tac2env.ml index 0fcdba1ca7..a36d022e4d 100644 --- a/tac2env.ml +++ b/tac2env.ml @@ -13,11 +13,6 @@ open Names open Libnames open Tac2expr -type ltac_constant = KerName.t -type ltac_constructor = KerName.t -type ltac_projection = KerName.t -type type_constant = KerName.t - type constructor_data = { cdata_prms : int; cdata_type : type_constant; @@ -121,7 +116,7 @@ struct id, (DirPath.repr dir) end -type tacref = +type tacref = Tac2expr.tacref = | TacConstant of ltac_constant | TacConstructor of ltac_constructor diff --git a/tac2env.mli b/tac2env.mli index 477f4ebec7..c4b8c1e0ca 100644 --- a/tac2env.mli +++ b/tac2env.mli @@ -63,10 +63,6 @@ val interp_projection : ltac_projection -> projection_data (** {5 Name management} *) -type tacref = -| TacConstant of ltac_constant -| TacConstructor of ltac_constructor - val push_ltac : visibility -> full_path -> tacref -> unit val locate_ltac : qualid -> tacref val locate_extended_all_ltac : qualid -> tacref list diff --git a/tac2expr.mli b/tac2expr.mli index 1840b567b4..63207ac78f 100644 --- a/tac2expr.mli +++ b/tac2expr.mli @@ -22,6 +22,10 @@ type ltac_constructor = KerName.t type ltac_projection = KerName.t type type_constant = KerName.t +type 'a or_relid = +| RelId of qualid located +| AbsKn of 'a + (** {5 Misc} *) type ml_tactic_name = { @@ -35,7 +39,7 @@ type raw_typexpr = | CTypVar of Name.t located | CTypArrow of Loc.t * raw_typexpr * raw_typexpr | CTypTuple of Loc.t * raw_typexpr list -| CTypRef of Loc.t * qualid located * raw_typexpr list +| CTypRef of Loc.t * type_constant or_relid * raw_typexpr list type raw_typedef = | CTydDef of raw_typexpr option @@ -66,15 +70,19 @@ type atom = | AtmInt of int | AtmStr of string +type tacref = +| TacConstant of ltac_constant +| TacConstructor of ltac_constructor + (** Tactic expressions *) type raw_patexpr = | CPatAny of Loc.t -| CPatRef of Loc.t * qualid located * raw_patexpr list +| CPatRef of Loc.t * ltac_constructor or_relid * raw_patexpr list | CPatTup of raw_patexpr list located type raw_tacexpr = | CTacAtm of atom located -| CTacRef of qualid located +| CTacRef of tacref or_relid | CTacFun of Loc.t * (Name.t located * raw_typexpr option) list * raw_tacexpr | CTacApp of Loc.t * raw_tacexpr * raw_tacexpr list | CTacLet of Loc.t * rec_flag * (Name.t located * raw_typexpr option * raw_tacexpr) list * raw_tacexpr @@ -85,13 +93,13 @@ type raw_tacexpr = | CTacSeq of Loc.t * raw_tacexpr * raw_tacexpr | CTacCse of Loc.t * raw_tacexpr * raw_taccase list | CTacRec of Loc.t * raw_recexpr -| CTacPrj of Loc.t * raw_tacexpr * qualid located -| CTacSet of Loc.t * raw_tacexpr * qualid located * raw_tacexpr +| CTacPrj of Loc.t * raw_tacexpr * ltac_projection or_relid +| CTacSet of Loc.t * raw_tacexpr * ltac_projection or_relid * raw_tacexpr | CTacExt of Loc.t * raw_generic_argument and raw_taccase = raw_patexpr * raw_tacexpr -and raw_recexpr = (qualid located * raw_tacexpr) list +and raw_recexpr = (ltac_projection or_relid * raw_tacexpr) list type case_info = | GCaseTuple of int diff --git a/tac2intern.ml b/tac2intern.ml index 32d1e07c17..c7e02a7b06 100644 --- a/tac2intern.ml +++ b/tac2intern.ml @@ -190,7 +190,8 @@ let push_name id t env = match id with let loc_of_tacexpr = function | CTacAtm (loc, _) -> loc -| CTacRef (loc, _) -> loc +| CTacRef (RelId (loc, _)) -> loc +| CTacRef (AbsKn _) -> Loc.ghost | CTacFun (loc, _, _) -> loc | CTacApp (loc, _, _) -> loc | CTacLet (loc, _, _, _) -> loc @@ -230,9 +231,10 @@ let rec subst_type subst (t : 'a glb_typexpr) = match t with let rec intern_type env (t : raw_typexpr) : UF.elt glb_typexpr = match t with | CTypVar (loc, Name id) -> GTypVar (get_alias (loc, id) env) | CTypVar (_, Anonymous) -> GTypVar (fresh_id env) -| CTypRef (_, (loc, qid), args) -> - let (dp, id) = repr_qualid qid in - let (kn, nparams) = +| CTypRef (loc, rel, args) -> + let (kn, nparams) = match rel with + | RelId (loc, qid) -> + let (dp, id) = repr_qualid qid in if DirPath.is_empty dp && Id.Map.mem id env.env_rec then Id.Map.find id env.env_rec else @@ -243,10 +245,17 @@ let rec intern_type env (t : raw_typexpr) : UF.elt glb_typexpr = match t with in let (nparams, _) = Tac2env.interp_type kn in (kn, nparams) + | AbsKn kn -> + let (nparams, _) = Tac2env.interp_type kn in + (kn, nparams) in let nargs = List.length args in let () = if not (Int.equal nparams nargs) then + let loc, qid = match rel with + | RelId lid -> lid + | AbsKn kn -> loc, shortest_qualid_of_type kn + in user_err ~loc (strbrk "The type constructor " ++ pr_qualid qid ++ strbrk " expects " ++ int nparams ++ strbrk " argument(s), but is here \ applied to " ++ int nargs ++ strbrk "argument(s)") @@ -467,7 +476,8 @@ let check_redundant_clause = function | [] -> () | (p, _) :: _ -> warn_redundant_clause ~loc:(loc_of_patexpr p) () -let get_variable env (loc, qid) = +let get_variable env var = match var with +| RelId (loc, qid) -> let (dp, id) = repr_qualid qid in if DirPath.is_empty dp && Id.Map.mem id env.env_var then ArgVar (loc, id) else @@ -477,10 +487,12 @@ let get_variable env (loc, qid) = CErrors.user_err ~loc (str "Unbound value " ++ pr_qualid qid) in ArgArg kn +| AbsKn kn -> ArgArg kn -let get_constructor env (loc, qid) = +let get_constructor env var = match var with +| RelId (loc, qid) -> let c = try Some (Tac2env.locate_ltac qid) with Not_found -> None in - match c with + begin match c with | Some (TacConstructor knc) -> let kn = Tac2env.interp_constructor knc in ArgArg (kn, knc) @@ -491,12 +503,19 @@ let get_constructor env (loc, qid) = let (dp, id) = repr_qualid qid in if DirPath.is_empty dp then ArgVar (loc, id) else CErrors.user_err ~loc (str "Unbound constructor " ++ pr_qualid qid) + end +| AbsKn knc -> + let kn = Tac2env.interp_constructor knc in + ArgArg (kn, knc) -let get_projection (loc, qid) = +let get_projection var = match var with +| RelId (loc, qid) -> let kn = try Tac2env.locate_projection qid with Not_found -> user_err ~loc (pr_qualid qid ++ str " is not a projection") in Tac2env.interp_projection kn +| AbsKn kn -> + Tac2env.interp_projection kn let intern_atm env = function | AtmInt n -> (GTacAtm (AtmInt n), GTypRef (t_int, [])) @@ -526,8 +545,8 @@ let rec intern_patexpr env = function end | CPatRef (_, qid, pl) -> begin match get_constructor env qid with - | ArgVar (loc, _) -> - user_err ~loc (str "Unbound constructor " ++ pr_qualid (snd qid)) + | ArgVar (loc, id) -> + user_err ~loc (str "Unbound constructor " ++ Nameops.pr_id id) | ArgArg (_, kn) -> GPatRef (kn, List.map (fun p -> intern_patexpr env p) pl) end | CPatTup (_, pl) -> @@ -565,7 +584,7 @@ let is_constructor env qid = match get_variable env qid with let rec intern_rec env = function | CTacAtm (_, atm) -> intern_atm env atm -| CTacRef qid -> +| CTacRef qid as e -> begin match get_variable env qid with | ArgVar (_, id) -> let sch = Id.Map.find id env.env_var in @@ -574,7 +593,8 @@ let rec intern_rec env = function let (_, _, sch) = Tac2env.interp_global kn in (GTacRef kn, fresh_type_scheme env sch) | ArgArg (TacConstructor kn) -> - intern_constructor env (fst qid) kn [] + let loc = loc_of_tacexpr e in + intern_constructor env loc kn [] end | CTacFun (loc, bnd, e) -> let fold (env, bnd, tl) ((_, na), t) = @@ -590,12 +610,13 @@ let rec intern_rec env = function let (e, t) = intern_rec env e in let t = List.fold_left (fun accu t -> GTypArrow (t, accu)) t tl in (GTacFun (bnd, e), t) -| CTacApp (loc, CTacRef qid, args) when is_constructor env qid -> +| CTacApp (loc, CTacRef qid, args) as e when is_constructor env qid -> let kn = match get_variable env qid with | ArgArg (TacConstructor kn) -> kn | _ -> assert false in - intern_constructor env (fst qid) kn args + let loc = loc_of_tacexpr e in + intern_constructor env loc kn args | CTacApp (loc, f, args) -> let (f, ft) = intern_rec env f in let fold arg (args, t) = @@ -687,7 +708,10 @@ let rec intern_rec env = function let pinfo = get_projection proj in let () = if not pinfo.pdata_mutb then - let (loc, _) = proj in + let loc = match proj with + | RelId (loc, _) -> loc + | AbsKn _ -> Loc.ghost + in user_err ~loc (str "Field is not mutable") in let subst = Array.init pinfo.pdata_prms (fun _ -> fresh_id env) in @@ -974,8 +998,12 @@ and intern_constructor env loc kn args = error_nargs_mismatch loc nargs (List.length args) and intern_record env loc fs = - let map ((loc, qid), e) = - let proj = get_projection (loc, qid) in + let map (proj, e) = + let loc = match proj with + | RelId (loc, _) -> loc + | AbsKn _ -> Loc.ghost + in + let proj = get_projection proj in (loc, proj, e) in let fs = List.map map fs in -- cgit v1.2.3 From a16d9c10b874a38fd4901e7d946d975ad49592c5 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 15 May 2017 15:57:01 +0200 Subject: Introducing tactic notations in Ltac2. --- g_ltac2.ml4 | 31 +++++++- tac2core.ml | 74 +++++++++++++++++- tac2entries.ml | 146 +++++++++++++++++++++++++++++++++++- tac2entries.mli | 23 ++++++ tac2expr.mli | 21 +++++- tac2intern.ml | 228 +++++++++++++++++++++++++++++++++++++++++++++++++++++++- tac2intern.mli | 8 ++ tac2print.ml | 4 +- tac2print.mli | 9 +++ 9 files changed, 533 insertions(+), 11 deletions(-) diff --git a/g_ltac2.ml4 b/g_ltac2.ml4 index 30f5399a88..565be4a199 100644 --- a/g_ltac2.ml4 +++ b/g_ltac2.ml4 @@ -16,11 +16,12 @@ open Misctypes open Tac2expr open Ltac_plugin -let tac2expr = Gram.entry_create "tactic:tac2expr" +let tac2expr = Tac2entries.Pltac.tac2expr let tac2type = Gram.entry_create "tactic:tac2type" let tac2def_val = Gram.entry_create "tactic:tac2def_val" let tac2def_typ = Gram.entry_create "tactic:tac2def_typ" let tac2def_ext = Gram.entry_create "tactic:tac2def_ext" +let tac2def_syn = Gram.entry_create "tactic:tac2def_syn" let tac2mode = Gram.entry_create "vernac:ltac2_command" let inj_wit wit loc x = CTacExt (loc, Genarg.in_gen (Genarg.rawwit wit) x) @@ -29,7 +30,7 @@ let inj_open_constr loc c = inj_wit Stdarg.wit_open_constr loc c let inj_ident loc c = inj_wit Stdarg.wit_ident loc c GEXTEND Gram - GLOBAL: tac2expr tac2type tac2def_val tac2def_typ tac2def_ext; + GLOBAL: tac2expr tac2type tac2def_val tac2def_typ tac2def_ext tac2def_syn; tac2pat: [ "1" LEFTA [ id = Prim.qualid; pl = LIST1 tac2pat LEVEL "0" -> CPatRef (!@loc, RelId id, pl) @@ -194,6 +195,30 @@ GEXTEND Gram StrPrm (id, t, ml) ] ] ; + syn_node: + [ [ "_" -> (!@loc, None) + | id = Prim.ident -> (!@loc, Some id) + ] ] + ; + sexpr: + [ [ s = Prim.string -> SexprStr (!@loc, s) + | n = Prim.integer -> SexprInt (!@loc, n) + | id = syn_node -> SexprRec (!@loc, id, []) + | id = syn_node; "("; tok = LIST1 sexpr SEP "," ; ")" -> + SexprRec (!@loc, id, tok) + ] ] + ; + syn_level: + [ [ -> None + | ":"; n = Prim.integer -> Some n + ] ] + ; + tac2def_syn: + [ [ "Notation"; toks = LIST1 sexpr; n = syn_level; ":="; + e = tac2expr -> + StrSyn (toks, n, e) + ] ] + ; END GEXTEND Gram @@ -212,6 +237,7 @@ PRINTED BY pr_ltac2entry | [ tac2def_val(v) ] -> [ v ] | [ tac2def_typ(t) ] -> [ t ] | [ tac2def_ext(e) ] -> [ e ] +| [ tac2def_syn(e) ] -> [ e ] END VERNAC COMMAND EXTEND VernacDeclareTactic2Definition CLASSIFIED AS SIDEFF @@ -250,3 +276,4 @@ open Stdarg VERNAC COMMAND EXTEND Ltac2Print CLASSIFIED AS SIDEFF | [ "Print" "Ltac2" reference(tac) ] -> [ Tac2entries.print_ltac tac ] END + diff --git a/tac2core.ml b/tac2core.ml index 3232fcba5b..ad238e6b8f 100644 --- a/tac2core.ml +++ b/tac2core.ml @@ -48,10 +48,14 @@ let t_unit = coq_core "unit" let t_list = coq_core "list" let t_constr = coq_core "constr" let t_ident = coq_core "ident" +let t_option = coq_core "option" let c_nil = coq_core "[]" let c_cons = coq_core "::" +let c_none = coq_core "None" +let c_some = coq_core "Some" + end open Core @@ -464,7 +468,6 @@ let interp_constr flags ist (c, _) = end let () = - let open Pretyping in let interp ist c = interp_constr (constr_flags ()) ist c in let obj = { ml_type = t_constr; @@ -473,7 +476,6 @@ let () = define_ml_object Stdarg.wit_constr obj let () = - let open Pretyping in let interp ist c = interp_constr (open_constr_no_classes_flags ()) ist c in let obj = { ml_type = t_constr; @@ -502,3 +504,71 @@ let () = (EConstr.of_constr c, sigma) in Pretyping.register_constr_interp0 wit_ltac2 interp + +(** Built-in notation scopes *) + +let add_scope s f = + Tac2entries.register_scope (Id.of_string s) f + +let scope_fail () = CErrors.user_err (str "Invalid parsing token") + +let rthunk e = + let loc = Tac2intern.loc_of_tacexpr e in + let var = [(loc, Anonymous), Some (CTypRef (loc, AbsKn Core.t_unit, []))] in + CTacFun (loc, var, e) + +let add_generic_scope s entry arg = + let parse = function + | [] -> + let scope = Extend.Aentry entry in + let act x = rthunk (CTacExt (Loc.ghost, in_gen (rawwit arg) x)) in + Tac2entries.ScopeRule (scope, act) + | _ -> scope_fail () + in + add_scope s parse + +let () = add_scope "list0" begin function +| [tok] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let scope = Extend.Alist0 scope in + let act l = + let l = List.map act l in + CTacLst (Loc.ghost, l) + in + Tac2entries.ScopeRule (scope, act) +| [tok; SexprStr (_, str)] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let sep = Extend.Atoken (CLexer.terminal str) in + let scope = Extend.Alist0sep (scope, sep) in + let act l = + let l = List.map act l in + CTacLst (Loc.ghost, l) + in + Tac2entries.ScopeRule (scope, act) +| _ -> scope_fail () +end + +let () = add_scope "opt" begin function +| [tok] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let scope = Extend.Aopt scope in + let act opt = match opt with + | None -> + CTacRef (AbsKn (TacConstructor Core.c_none)) + | Some x -> + CTacApp (Loc.ghost, CTacRef (AbsKn (TacConstructor Core.c_some)), [act x]) + in + Tac2entries.ScopeRule (scope, act) +| _ -> scope_fail () +end + +let () = add_scope "self" begin function +| [] -> + let scope = Extend.Aself in + let act tac = rthunk tac in + Tac2entries.ScopeRule (scope, act) +| _ -> scope_fail () +end + +let () = add_generic_scope "ident" Pcoq.Prim.ident Stdarg.wit_ident +let () = add_generic_scope "constr" Pcoq.Constr.constr Stdarg.wit_constr diff --git a/tac2entries.ml b/tac2entries.ml index 374a367188..fbfc687ee7 100644 --- a/tac2entries.ml +++ b/tac2entries.ml @@ -19,6 +19,13 @@ open Tac2print open Tac2intern open Vernacexpr +(** Grammar entries *) + +module Pltac = +struct +let tac2expr = Pcoq.Gram.entry_create "tactic:tac2expr" +end + (** Tactic definition *) type tacdef = { @@ -98,8 +105,6 @@ let next i = let () = incr i in ans -let dummy_var i = Id.of_string (Printf.sprintf "_%i" i) - let define_typedef kn (params, def as qdef) = match def with | GTydDef _ -> Tac2env.define_type kn qdef @@ -435,10 +440,147 @@ let register_type ?local isrec types = match types with let types = List.map map types in register_typedef ?local isrec types +(** Parsing *) + +type 'a token = +| TacTerm of string +| TacNonTerm of Name.t * 'a + +type scope_rule = +| ScopeRule : (raw_tacexpr, 'a) Extend.symbol * ('a -> raw_tacexpr) -> scope_rule + +type scope_interpretation = sexpr list -> scope_rule + +let scope_table : scope_interpretation Id.Map.t ref = ref Id.Map.empty + +let register_scope id s = + scope_table := Id.Map.add id s !scope_table + +module ParseToken = +struct + +let loc_of_token = function +| SexprStr (loc, _) -> loc +| SexprInt (loc, _) -> loc +| SexprRec (loc, _, _) -> loc + +let parse_scope = function +| SexprRec (_, (loc, Some id), toks) -> + if Id.Map.mem id !scope_table then + Id.Map.find id !scope_table toks + else + CErrors.user_err ~loc (str "Unknown scope" ++ spc () ++ Nameops.pr_id id) +| tok -> + let loc = loc_of_token tok in + CErrors.user_err ~loc (str "Invalid parsing token") + +let parse_token = function +| SexprStr (_, s) -> TacTerm s +| SexprRec (_, (_, na), [tok]) -> + let na = match na with None -> Anonymous | Some id -> Name id in + let scope = parse_scope tok in + TacNonTerm (na, scope) +| tok -> + let loc = loc_of_token tok in + CErrors.user_err ~loc (str "Invalid parsing token") + +end + +let parse_scope = ParseToken.parse_scope + +type synext = { + synext_tok : sexpr list; + synext_exp : raw_tacexpr; + synext_lev : int option; + synext_loc : bool; +} + +type krule = +| KRule : + (raw_tacexpr, 'act, Loc.t -> raw_tacexpr) Extend.rule * + ((Loc.t -> (Name.t * raw_tacexpr) list -> raw_tacexpr) -> 'act) -> krule + +let rec get_rule (tok : scope_rule token list) : krule = match tok with +| [] -> KRule (Extend.Stop, fun k loc -> k loc []) +| TacNonTerm (na, ScopeRule (scope, inj)) :: tok -> + let KRule (rule, act) = get_rule tok in + let rule = Extend.Next (rule, scope) in + let act k e = act (fun loc acc -> k loc ((na, inj e) :: acc)) in + KRule (rule, act) +| TacTerm t :: tok -> + let KRule (rule, act) = get_rule tok in + let rule = Extend.Next (rule, Extend.Atoken (CLexer.terminal t)) in + let act k _ = act k in + KRule (rule, act) + +let perform_notation syn st = + let tok = List.rev_map ParseToken.parse_token syn.synext_tok in + let KRule (rule, act) = get_rule tok in + let mk loc args = + let map (na, e) = + let loc = loc_of_tacexpr e in + ((loc, na), None, e) + in + let bnd = List.map map args in + CTacLet (loc, false, bnd, syn.synext_exp) + in + let rule = Extend.Rule (rule, act mk) in + let lev = match syn.synext_lev with + | None -> None + | Some lev -> Some (string_of_int lev) + in + let rule = (lev, None, [rule]) in + ([Pcoq.ExtendRule (Pltac.tac2expr, None, (None, [rule]))], st) + +let ltac2_notation = + Pcoq.create_grammar_command "ltac2-notation" perform_notation + +let cache_synext (_, syn) = + Pcoq.extend_grammar_command ltac2_notation syn + +let open_synext i (_, syn) = + if Int.equal i 1 then Pcoq.extend_grammar_command ltac2_notation syn + +let subst_synext (subst, syn) = + let e = Tac2intern.subst_rawexpr subst syn.synext_exp in + if e == syn.synext_exp then syn else { syn with synext_exp = e } + +let classify_synext o = + if o.synext_loc then Dispose else Substitute o + +let inTac2Notation : synext -> obj = + declare_object {(default_object "TAC2-NOTATION") with + cache_function = cache_synext; + open_function = open_synext; + subst_function = subst_synext; + classify_function = classify_synext} + +let register_notation ?(local = false) tkn lev body = + (** Check that the tokens make sense *) + let entries = List.map ParseToken.parse_token tkn in + let fold accu tok = match tok with + | TacTerm _ -> accu + | TacNonTerm (Name id, _) -> Id.Set.add id accu + | TacNonTerm (Anonymous, _) -> accu + in + let ids = List.fold_left fold Id.Set.empty entries in + (** Globalize so that names are absolute *) + let body = Tac2intern.globalize ids body in + let ext = { + synext_tok = tkn; + synext_exp = body; + synext_lev = lev; + synext_loc = local; + } in + Lib.add_anonymous_leaf (inTac2Notation ext) + +(** Toplevel entries *) + let register_struct ?local str = match str with | StrVal (isrec, e) -> register_ltac ?local isrec e | StrTyp (isrec, t) -> register_type ?local isrec t | StrPrm (id, t, ml) -> register_primitive ?local id t ml +| StrSyn (tok, lev, e) -> register_notation ?local tok lev e (** Printing *) diff --git a/tac2entries.mli b/tac2entries.mli index 0d9b3ad134..71e8150057 100644 --- a/tac2entries.mli +++ b/tac2entries.mli @@ -24,6 +24,22 @@ val register_primitive : ?local:bool -> val register_struct : ?local:bool -> strexpr -> unit +val register_notation : ?local:bool -> sexpr list -> int option -> + raw_tacexpr -> unit + +(** {5 Notations} *) + +type scope_rule = +| ScopeRule : (raw_tacexpr, 'a) Extend.symbol * ('a -> raw_tacexpr) -> scope_rule + +type scope_interpretation = sexpr list -> scope_rule + +val register_scope : Id.t -> scope_interpretation -> unit +(** Create a new scope with the provided name *) + +val parse_scope : sexpr -> scope_rule +(** Use this to interpret the subscopes for interpretation functions *) + (** {5 Inspecting} *) val print_ltac : Libnames.reference -> unit @@ -32,3 +48,10 @@ val print_ltac : Libnames.reference -> unit (** Evaluate a tactic expression in the current environment *) val call : default:bool -> raw_tacexpr -> unit + +(** {5 Parsing entries} *) + +module Pltac : +sig +val tac2expr : raw_tacexpr Pcoq.Gram.entry +end diff --git a/tac2expr.mli b/tac2expr.mli index 63207ac78f..acdad9bab4 100644 --- a/tac2expr.mli +++ b/tac2expr.mli @@ -129,7 +129,23 @@ type glb_tacexpr = | GTacExt of glob_generic_argument | GTacPrm of ml_tactic_name * glb_tacexpr list -(** Toplevel statements *) +(** {5 Parsing & Printing} *) + +type exp_level = +| E5 +| E4 +| E3 +| E2 +| E1 +| E0 + +type sexpr = +| SexprStr of string located +| SexprInt of int located +| SexprRec of Loc.t * Id.t option located * sexpr list + +(** {5 Toplevel statements} *) + type strexpr = | StrVal of rec_flag * (Name.t located * raw_tacexpr) list (** Term definition *) @@ -137,6 +153,9 @@ type strexpr = (** Type definition *) | StrPrm of Id.t located * raw_typexpr * ml_tactic_name (** External definition *) +| StrSyn of sexpr list * int option * raw_tacexpr + (** Syntactic extensions *) + (** {5 Dynamic semantics} *) diff --git a/tac2intern.ml b/tac2intern.ml index c7e02a7b06..17d08b2285 100644 --- a/tac2intern.ml +++ b/tac2intern.ml @@ -476,10 +476,10 @@ let check_redundant_clause = function | [] -> () | (p, _) :: _ -> warn_redundant_clause ~loc:(loc_of_patexpr p) () -let get_variable env var = match var with +let get_variable0 mem var = match var with | RelId (loc, qid) -> let (dp, id) = repr_qualid qid in - if DirPath.is_empty dp && Id.Map.mem id env.env_var then ArgVar (loc, id) + if DirPath.is_empty dp && mem id then ArgVar (loc, id) else let kn = try Tac2env.locate_ltac qid @@ -489,6 +489,10 @@ let get_variable env var = match var with ArgArg kn | AbsKn kn -> ArgArg kn +let get_variable env var = + let mem id = Id.Map.mem id env.env_var in + get_variable0 mem var + let get_constructor env var = match var with | RelId (loc, qid) -> let c = try Some (Tac2env.locate_ltac qid) with Not_found -> None in @@ -1106,6 +1110,104 @@ let intern_open_type t = let t = normalize env (count, vars) t in (!count, t) +(** Globalization *) + +let add_name accu = function +| Name id -> Id.Set.add id accu +| Anonymous -> accu + +let get_projection0 var = match var with +| RelId (loc, qid) -> + let kn = try Tac2env.locate_projection qid with Not_found -> + user_err ~loc (pr_qualid qid ++ str " is not a projection") + in + kn +| AbsKn kn -> kn + +let rec globalize ids e = match e with +| CTacAtm _ -> e +| CTacRef ref -> + let mem id = Id.Set.mem id ids in + begin match get_variable0 mem ref with + | ArgVar _ -> e + | ArgArg kn -> CTacRef (AbsKn kn) + end +| CTacFun (loc, bnd, e) -> + let fold accu ((_, na), _) = add_name accu na in + let ids = List.fold_left fold ids bnd in + let e = globalize ids e in + CTacFun (loc, bnd, e) +| CTacApp (loc, e, el) -> + let e = globalize ids e in + let el = List.map (fun e -> globalize ids e) el in + CTacApp (loc, e, el) +| CTacLet (loc, isrec, bnd, e) -> + let fold accu ((_, na), _, _) = add_name accu na in + let ext = List.fold_left fold Id.Set.empty bnd in + let eids = Id.Set.union ext ids in + let e = globalize eids e in + let map (qid, t, e) = + let ids = if isrec then eids else ids in + (qid, t, globalize ids e) + in + let bnd = List.map map bnd in + CTacLet (loc, isrec, bnd, e) +| CTacTup (loc, el) -> + let el = List.map (fun e -> globalize ids e) el in + CTacTup (loc, el) +| CTacArr (loc, el) -> + let el = List.map (fun e -> globalize ids e) el in + CTacArr (loc, el) +| CTacLst (loc, el) -> + let el = List.map (fun e -> globalize ids e) el in + CTacLst (loc, el) +| CTacCnv (loc, e, t) -> + let e = globalize ids e in + CTacCnv (loc, e, t) +| CTacSeq (loc, e1, e2) -> + let e1 = globalize ids e1 in + let e2 = globalize ids e2 in + CTacSeq (loc, e1, e2) +| CTacCse (loc, e, bl) -> + let e = globalize ids e in + let bl = List.map (fun b -> globalize_case ids b) bl in + CTacCse (loc, e, bl) +| CTacRec (loc, r) -> + let map (p, e) = + let p = get_projection0 p in + let e = globalize ids e in + (AbsKn p, e) + in + CTacRec (loc, List.map map r) +| CTacPrj (loc, e, p) -> + let e = globalize ids e in + let p = get_projection0 p in + CTacPrj (loc, e, AbsKn p) +| CTacSet (loc, e, p, e') -> + let e = globalize ids e in + let p = get_projection0 p in + let e' = globalize ids e' in + CTacSet (loc, e, AbsKn p, e') +| CTacExt (loc, arg) -> + let arg = pr_argument_type (genarg_tag arg) in + CErrors.user_err ~loc (str "Cannot globalize generic arguments of type" ++ spc () ++ arg) + +and globalize_case ids (p, e) = + (globalize_pattern ids p, globalize ids e) + +and globalize_pattern ids p = match p with +| CPatAny _ -> p +| CPatRef (loc, cst, pl) -> + let cst = match get_constructor () cst with + | ArgVar _ -> cst + | ArgArg (_, knc) -> AbsKn knc + in + let pl = List.map (fun p -> globalize_pattern ids p) pl in + CPatRef (loc, cst, pl) +| CPatTup (loc, pl) -> + let pl = List.map (fun p -> globalize_pattern ids p) pl in + CPatTup (loc, pl) + (** Kernel substitution *) open Mod_subst @@ -1213,6 +1315,128 @@ let subst_type_scheme subst (prm, t as sch) = let t' = subst_type subst t in if t' == t then sch else (prm, t') +let subst_or_relid subst ref = match ref with +| RelId _ -> ref +| AbsKn kn -> + let kn' = subst_kn subst kn in + if kn' == kn then ref else AbsKn kn' + +let rec subst_rawtype subst t = match t with +| CTypVar _ -> t +| CTypArrow (loc, t1, t2) -> + let t1' = subst_rawtype subst t1 in + let t2' = subst_rawtype subst t2 in + if t1' == t1 && t2' == t2 then t else CTypArrow (loc, t1', t2') +| CTypTuple (loc, tl) -> + let tl' = List.smartmap (fun t -> subst_rawtype subst t) tl in + if tl' == tl then t else CTypTuple (loc, tl') +| CTypRef (loc, ref, tl) -> + let ref' = subst_or_relid subst ref in + let tl' = List.smartmap (fun t -> subst_rawtype subst t) tl in + if ref' == ref && tl' == tl then t else CTypRef (loc, ref', tl') + +let subst_tacref subst ref = match ref with +| RelId _ -> ref +| AbsKn (TacConstant kn) -> + let kn' = subst_kn subst kn in + if kn' == kn then ref else AbsKn (TacConstant kn') +| AbsKn (TacConstructor kn) -> + let kn' = subst_kn subst kn in + if kn' == kn then ref else AbsKn (TacConstructor kn') + +let subst_projection subst prj = match prj with +| RelId _ -> prj +| AbsKn kn -> + let kn' = subst_kn subst kn in + if kn' == kn then prj else AbsKn kn' + +let rec subst_rawpattern subst p = match p with +| CPatAny _ -> p +| CPatRef (loc, c, pl) -> + let pl' = List.smartmap (fun p -> subst_rawpattern subst p) pl in + let c' = match c with + | RelId _ -> c + | AbsKn kn -> + let kn' = subst_kn subst kn in + if kn' == kn then c else AbsKn kn' + in + if pl' == pl && c' == c then p else CPatRef (loc, c', pl') +| CPatTup (loc, pl) -> + let pl' = List.smartmap (fun p -> subst_rawpattern subst p) pl in + if pl' == pl then p else CPatTup (loc, pl') + +(** Used for notations *) +let rec subst_rawexpr subst t = match t with +| CTacAtm _ -> t +| CTacRef ref -> + let ref' = subst_tacref subst ref in + if ref' == ref then t else CTacRef ref' +| CTacFun (loc, bnd, e) -> + let map (na, t as p) = + let t' = Option.smartmap (fun t -> subst_rawtype subst t) t in + if t' == t then p else (na, t') + in + let bnd' = List.smartmap map bnd in + let e' = subst_rawexpr subst e in + if bnd' == bnd && e' == e then t else CTacFun (loc, bnd', e') +| CTacApp (loc, e, el) -> + let e' = subst_rawexpr subst e in + let el' = List.smartmap (fun e -> subst_rawexpr subst e) el in + if e' == e && el' == el then t else CTacApp (loc, e', el') +| CTacLet (loc, isrec, bnd, e) -> + let map (na, t, e as p) = + let t' = Option.smartmap (fun t -> subst_rawtype subst t) t in + let e' = subst_rawexpr subst e in + if t' == t && e' == e then p else (na, t', e') + in + let bnd' = List.smartmap map bnd in + let e' = subst_rawexpr subst e in + if bnd' == bnd && e' == e then t else CTacLet (loc, isrec, bnd', e') +| CTacTup (loc, el) -> + let el' = List.smartmap (fun e -> subst_rawexpr subst e) el in + if el' == el then t else CTacTup (loc, el') +| CTacArr (loc, el) -> + let el' = List.smartmap (fun e -> subst_rawexpr subst e) el in + if el' == el then t else CTacArr (loc, el') +| CTacLst (loc, el) -> + let el' = List.smartmap (fun e -> subst_rawexpr subst e) el in + if el' == el then t else CTacLst (loc, el') +| CTacCnv (loc, e, c) -> + let e' = subst_rawexpr subst e in + let c' = subst_rawtype subst c in + if c' == c && e' == e then t else CTacCnv (loc, e', c') +| CTacSeq (loc, e1, e2) -> + let e1' = subst_rawexpr subst e1 in + let e2' = subst_rawexpr subst e2 in + if e1' == e1 && e2' == e2 then t else CTacSeq (loc, e1', e2') +| CTacCse (loc, e, bl) -> + let map (p, e as x) = + let p' = subst_rawpattern subst p in + let e' = subst_rawexpr subst e in + if p' == p && e' == e then x else (p', e') + in + let e' = subst_rawexpr subst e in + let bl' = List.smartmap map bl in + if e' == e && bl' == bl then t else CTacCse (loc, e', bl') +| CTacRec (loc, el) -> + let map (prj, e as p) = + let prj' = subst_projection subst prj in + let e' = subst_rawexpr subst e in + if prj' == prj && e' == e then p else (prj', e') + in + let el' = List.smartmap map el in + if el' == el then t else CTacRec (loc, el') +| CTacPrj (loc, e, prj) -> + let prj' = subst_projection subst prj in + let e' = subst_rawexpr subst e in + if prj' == prj && e' == e then t else CTacPrj (loc, e', prj') +| CTacSet (loc, e, prj, r) -> + let prj' = subst_projection subst prj in + let e' = subst_rawexpr subst e in + let r' = subst_rawexpr subst r in + if prj' == prj && e' == e && r' == r then t else CTacSet (loc, e', prj', r') +| CTacExt _ -> assert false (** Should not be generated by gloabalization *) + (** Registering *) let () = diff --git a/tac2intern.mli b/tac2intern.mli index 4c482d0b0c..3d400a5cdd 100644 --- a/tac2intern.mli +++ b/tac2intern.mli @@ -27,6 +27,14 @@ val subst_expr : substitution -> glb_tacexpr -> glb_tacexpr val subst_quant_typedef : substitution -> glb_quant_typedef -> glb_quant_typedef val subst_type_scheme : substitution -> type_scheme -> type_scheme +val subst_rawexpr : substitution -> raw_tacexpr -> raw_tacexpr + +(** {5 Notations} *) + +val globalize : Id.Set.t -> raw_tacexpr -> raw_tacexpr +(** Replaces all qualified identifiers by their corresponding kernel name. The + set represents bound variables in the context. *) + (** Errors *) val error_nargs_mismatch : Loc.t -> int -> int -> 'a diff --git a/tac2print.ml b/tac2print.ml index ffa5ddc05a..e6f0582e3d 100644 --- a/tac2print.ml +++ b/tac2print.ml @@ -33,7 +33,7 @@ type typ_level = let pr_typref kn = Libnames.pr_qualid (Tac2env.shortest_qualid_of_type kn) -let rec pr_glbtype_gen pr lvl c = +let pr_glbtype_gen pr lvl c = let rec pr_glbtype lvl = function | GTypVar n -> str "'" ++ str (pr n) | GTypRef (kn, []) -> pr_typref kn @@ -88,7 +88,7 @@ let pr_constructor kn = let pr_projection kn = Libnames.pr_qualid (Tac2env.shortest_qualid_of_projection kn) -type exp_level = +type exp_level = Tac2expr.exp_level = | E5 | E4 | E3 diff --git a/tac2print.mli b/tac2print.mli index 94555a4c95..ddd599641d 100644 --- a/tac2print.mli +++ b/tac2print.mli @@ -12,13 +12,22 @@ open Tac2expr (** {5 Printing types} *) +type typ_level = +| T5_l +| T5_r +| T2 +| T1 +| T0 + val pr_typref : type_constant -> std_ppcmds +val pr_glbtype_gen : ('a -> string) -> typ_level -> 'a glb_typexpr -> std_ppcmds val pr_glbtype : ('a -> string) -> 'a glb_typexpr -> std_ppcmds (** {5 Printing expressions} *) val pr_constructor : ltac_constructor -> std_ppcmds val pr_projection : ltac_projection -> std_ppcmds +val pr_glbexpr_gen : exp_level -> glb_tacexpr -> std_ppcmds val pr_glbexpr : glb_tacexpr -> std_ppcmds (** {5 Utilities} *) -- cgit v1.2.3 From df1c50b36f4927fdf64a3ed99a4a077f5175ac5e Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 16 May 2017 18:03:55 +0200 Subject: Removing dead code in Ltac2, and cleaning up a bit. --- g_ltac2.ml4 | 1 - tac2core.ml | 12 ------------ tac2core.mli | 9 +++++++++ tac2entries.ml | 1 - tac2env.ml | 1 - tac2intern.ml | 3 --- 6 files changed, 9 insertions(+), 18 deletions(-) diff --git a/g_ltac2.ml4 b/g_ltac2.ml4 index 565be4a199..51919addf2 100644 --- a/g_ltac2.ml4 +++ b/g_ltac2.ml4 @@ -8,7 +8,6 @@ open Pp open Util -open Genarg open Names open Pcoq open Constrexpr diff --git a/tac2core.ml b/tac2core.ml index ad238e6b8f..13395c87b3 100644 --- a/tac2core.ml +++ b/tac2core.ml @@ -8,7 +8,6 @@ open CSig open Pp -open CErrors open Names open Genarg open Geninterp @@ -20,10 +19,6 @@ open Proofview.Notations (** Standard values *) let coq_core n = KerName.make2 Tac2env.coq_prefix (Label.of_id (Id.of_string_soft n)) -let stdlib_prefix md = - MPfile (DirPath.make (List.map Id.of_string [md; "ltac2"; "Coq"])) -let coq_stdlib md n = - KerName.make2 (stdlib_prefix md) (Label.of_id (Id.of_string n)) let val_tag t = match val_tag t with | Val.Base t -> t @@ -141,9 +136,6 @@ let err_notfocussed = let err_outofbounds = LtacError (coq_core "Out_of_bounds", [||]) -let err_notfound = - LtacError (coq_core "Not_found", [||]) - (** Helper functions *) let thaw f = interp_app f [v_unit] @@ -158,10 +150,6 @@ let wrap f = let wrap_unit f = return () >>= fun () -> f (); return v_unit -let wrap_exn f err = - return () >>= fun () -> - try return (f ()) with e when CErrors.noncritical e -> err - let pf_apply f = Proofview.Goal.goals >>= function | [] -> diff --git a/tac2core.mli b/tac2core.mli index 27144bc6e2..fc90499ac6 100644 --- a/tac2core.mli +++ b/tac2core.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open Names open Tac2env open Tac2expr @@ -18,6 +19,11 @@ val t_list : type_constant val c_nil : ltac_constant val c_cons : ltac_constant +val t_int : type_constant +val t_option : type_constant +val t_string : type_constant +val t_array : type_constant + end (** {5 Ltac2 FFI} *) @@ -50,4 +56,7 @@ val to_constr : valexpr -> EConstr.t val of_exn : Exninfo.iexn -> valexpr val to_exn : valexpr -> Exninfo.iexn +val of_ident : Id.t -> valexpr +val to_ident : valexpr -> Id.t + end diff --git a/tac2entries.ml b/tac2entries.ml index fbfc687ee7..3959e705ed 100644 --- a/tac2entries.ml +++ b/tac2entries.ml @@ -13,7 +13,6 @@ open Names open Libnames open Libobject open Nametab -open Tac2env open Tac2expr open Tac2print open Tac2intern diff --git a/tac2env.ml b/tac2env.ml index a36d022e4d..5ccdd018ee 100644 --- a/tac2env.ml +++ b/tac2env.ml @@ -8,7 +8,6 @@ open CErrors open Util -open Genarg open Names open Libnames open Tac2expr diff --git a/tac2intern.ml b/tac2intern.ml index 17d08b2285..756bbe3076 100644 --- a/tac2intern.ml +++ b/tac2intern.ml @@ -39,7 +39,6 @@ type 'a t val equal : elt -> elt -> bool val create : unit -> 'a t val fresh : 'a t -> elt -val size : 'a t -> int val find : elt -> 'a t -> (elt * 'a option) val union : elt -> elt -> 'a t -> unit val set : elt -> 'a -> 'a t -> unit @@ -69,8 +68,6 @@ type 'a t = { mutable uf_size : int; } -let size p = p.uf_size - let resize p = if Int.equal (Array.length p.uf_data) p.uf_size then begin let nsize = 2 * p.uf_size + 1 in -- cgit v1.2.3 From 6a4d15c6ce3994509085ef575cc2f242925af15a Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 16 May 2017 18:14:40 +0200 Subject: Extending the Coq API in Ltac2. --- Constr.v | 34 ++++++++++++++++++++++++++++++++++ Control.v | 5 +++++ Init.v | 12 +++++++++++- tac2core.ml | 41 +++++++++++++++++++++++++++++++++++++++++ 4 files changed, 91 insertions(+), 1 deletion(-) diff --git a/Constr.v b/Constr.v index 994f9bf3ac..9e5833a778 100644 --- a/Constr.v +++ b/Constr.v @@ -7,3 +7,37 @@ (************************************************************************) Require Import Coq.ltac2.Init. + +Ltac2 @ external type : constr -> constr := "ltac2" "constr_type". +(** Return the type of a term *) + +Ltac2 @ external equal : constr -> constr -> bool := "ltac2" "constr_equal". +(** Strict syntactic equality: only up to α-conversion and evar expansion *) + +Module Unsafe. + +(** Low-level access to kernel term. Use with care! *) + +Ltac2 Type kind := [ +| Rel (int) +| Var (ident) +| Meta (meta) +| Evar (evar, constr list) +| Sort (sort) +| Cast (constr, cast, constr) +| Prod (ident option, constr, constr) +| Lambda (ident option, constr, constr) +| LetIn (ident option, constr, constr, constr) +| App (constr, constr list) +| Constant (constant, instance) +| Ind (inductive, instance) +| Constructor (inductive, instance) +(* + | Case of case_info * 'constr * 'constr * 'constr array + | Fix of ('constr, 'types) pfixpoint + | CoFix of ('constr, 'types) pcofixpoint +*) +| Proj (projection, constr) +]. + +End Unsafe. diff --git a/Control.v b/Control.v index a476513ede..6b3b360abb 100644 --- a/Control.v +++ b/Control.v @@ -28,6 +28,11 @@ Ltac2 @ external focus : int -> int -> (unit -> 'a) -> 'a := "ltac2" "focus". Ltac2 @ external shelve : unit -> unit := "ltac2" "shelve". Ltac2 @ external shelve_unifiable : unit -> unit := "ltac2" "shelve_unifiable". +Ltac2 @ external new_goal : evar -> unit := "ltac2" "new_goal". +(** Adds the given evar to the list of goals as the last one. If it is + already defined in the current state, don't do anything. Panics if the + evar is not in the current state. *) + (** Goal inspection *) Ltac2 @ external goal : unit -> constr := "ltac2" "goal". diff --git a/Init.v b/Init.v index 8ff5837bb4..1d2d40f5c0 100644 --- a/Init.v +++ b/Init.v @@ -15,10 +15,20 @@ Global Set Default Proof Mode "Ltac2". Ltac2 Type int. Ltac2 Type string. Ltac2 Type char. +Ltac2 Type ident. +(** Constr-specific built-in types *) +Ltac2 Type meta. Ltac2 Type evar. +Ltac2 Type sort. +Ltac2 Type cast. +Ltac2 Type instance. +Ltac2 Type constant. +Ltac2 Type inductive. +Ltac2 Type constructor. +Ltac2 Type projection. Ltac2 Type constr. -Ltac2 Type ident. + Ltac2 Type message. Ltac2 Type exn := [ .. ]. Ltac2 Type 'a array. diff --git a/tac2core.ml b/tac2core.ml index 13395c87b3..94758eb217 100644 --- a/tac2core.ml +++ b/tac2core.ml @@ -136,6 +136,9 @@ let err_notfocussed = let err_outofbounds = LtacError (coq_core "Out_of_bounds", [||]) +let err_notfound = + LtacError (coq_core "Not_found", [||]) + (** Helper functions *) let thaw f = interp_app f [v_unit] @@ -270,6 +273,31 @@ let prm_string_get : ml_tactic = function else wrap (fun () -> Value.of_char (Bytes.get s n)) | _ -> assert false +(** Terms *) + +(** constr -> constr *) +let prm_constr_type : ml_tactic = function +| [c] -> + let c = Value.to_constr c in + let get_type env sigma = + Proofview.V82.wrap_exceptions begin fun () -> + let (sigma, t) = Typing.type_of env sigma c in + let t = Value.of_constr t in + Proofview.Unsafe.tclEVARS sigma <*> Proofview.tclUNIT t + end in + pf_apply get_type +| _ -> assert false + +(** constr -> constr *) +let prm_constr_equal : ml_tactic = function +| [c1; c2] -> + let c1 = Value.to_constr c1 in + let c2 = Value.to_constr c2 in + Proofview.tclEVARMAP >>= fun sigma -> + let b = EConstr.eq_constr sigma c1 c2 in + Proofview.tclUNIT (Value.of_bool b) +| _ -> assert false + (** Error *) let prm_throw : ml_tactic = function @@ -342,6 +370,15 @@ let prm_shelve_unifiable : ml_tactic = function | [_] -> Proofview.shelve_unifiable >>= fun () -> return v_unit | _ -> assert false +let prm_new_goal : ml_tactic = function +| [ev] -> + let ev = Evar.unsafe_of_int (Value.to_int ev) in + Proofview.tclEVARMAP >>= fun sigma -> + if Evd.mem sigma ev then + Proofview.Unsafe.tclNEWGOALS [ev] <*> Proofview.tclUNIT v_unit + else throw err_notfound +| _ -> assert false + (** unit -> constr *) let prm_goal : ml_tactic = function | [_] -> @@ -391,6 +428,9 @@ let () = Tac2env.define_primitive (pname "string_length") prm_string_length let () = Tac2env.define_primitive (pname "string_get") prm_string_get let () = Tac2env.define_primitive (pname "string_set") prm_string_set +let () = Tac2env.define_primitive (pname "constr_type") prm_constr_type +let () = Tac2env.define_primitive (pname "constr_equal") prm_constr_equal + let () = Tac2env.define_primitive (pname "int_equal") prm_int_equal let () = Tac2env.define_primitive (pname "int_compare") prm_int_compare let () = Tac2env.define_primitive (pname "int_neg") prm_int_neg @@ -410,6 +450,7 @@ let () = Tac2env.define_primitive (pname "enter") prm_enter let () = Tac2env.define_primitive (pname "focus") prm_focus let () = Tac2env.define_primitive (pname "shelve") prm_shelve let () = Tac2env.define_primitive (pname "shelve_unifiable") prm_shelve_unifiable +let () = Tac2env.define_primitive (pname "new_goal") prm_new_goal let () = Tac2env.define_primitive (pname "goal") prm_goal let () = Tac2env.define_primitive (pname "hyp") prm_hyp let () = Tac2env.define_primitive (pname "refine") prm_refine -- cgit v1.2.3 From 0bfa6d3cda461f4d09ec0bfa9781042199b1f43b Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 19 May 2017 15:25:58 +0200 Subject: Adding new tactic notation scopes. --- tac2core.ml | 43 +++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 43 insertions(+) diff --git a/tac2core.ml b/tac2core.ml index 94758eb217..c82893efc2 100644 --- a/tac2core.ml +++ b/tac2core.ml @@ -577,6 +577,27 @@ let () = add_scope "list0" begin function | _ -> scope_fail () end +let () = add_scope "list1" begin function +| [tok] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let scope = Extend.Alist1 scope in + let act l = + let l = List.map act l in + CTacLst (Loc.ghost, l) + in + Tac2entries.ScopeRule (scope, act) +| [tok; SexprStr (_, str)] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let sep = Extend.Atoken (CLexer.terminal str) in + let scope = Extend.Alist1sep (scope, sep) in + let act l = + let l = List.map act l in + CTacLst (Loc.ghost, l) + in + Tac2entries.ScopeRule (scope, act) +| _ -> scope_fail () +end + let () = add_scope "opt" begin function | [tok] -> let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in @@ -599,5 +620,27 @@ let () = add_scope "self" begin function | _ -> scope_fail () end +let () = add_scope "next" begin function +| [] -> + let scope = Extend.Anext in + let act tac = rthunk tac in + Tac2entries.ScopeRule (scope, act) +| _ -> scope_fail () +end + +let () = add_scope "tactic" begin function +| [] -> + (** Default to level 5 parsing *) + let scope = Extend.Aentryl (Tac2entries.Pltac.tac2expr, 5) in + let act tac = rthunk tac in + Tac2entries.ScopeRule (scope, act) +| [SexprInt (loc, n)] -> + let () = if n < 0 || n > 5 then scope_fail () in + let scope = Extend.Aentryl (Tac2entries.Pltac.tac2expr, n) in + let act tac = rthunk tac in + Tac2entries.ScopeRule (scope, act) +| _ -> scope_fail () +end + let () = add_generic_scope "ident" Pcoq.Prim.ident Stdarg.wit_ident let () = add_generic_scope "constr" Pcoq.Constr.constr Stdarg.wit_constr -- cgit v1.2.3 From c25e86e6b4e8bb694d3c8e50f04a92cc33ad807d Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 24 Jul 2017 17:41:01 +0200 Subject: Turning the ltac2 subfolder into a standalone plugin. --- .gitignore | 14 + Array.v | 14 - Constr.v | 43 -- Control.v | 49 -- Init.v | 56 -- Int.v | 18 - LICENSE | 458 +++++++++++++++ Ltac2.v | 16 - Makefile | 12 + Message.v | 20 - README.md | 0 String.v | 14 - _CoqProject | 28 + g_ltac2.ml4 | 278 --------- ltac2_plugin.mlpack | 7 - src/g_ltac2.ml4 | 278 +++++++++ src/ltac2_plugin.mlpack | 7 + src/tac2core.ml | 648 +++++++++++++++++++++ src/tac2core.mli | 62 ++ src/tac2entries.ml | 648 +++++++++++++++++++++ src/tac2entries.mli | 57 ++ src/tac2env.ml | 242 ++++++++ src/tac2env.mli | 106 ++++ src/tac2expr.mli | 195 +++++++ src/tac2intern.ml | 1454 +++++++++++++++++++++++++++++++++++++++++++++++ src/tac2intern.mli | 41 ++ src/tac2interp.ml | 160 ++++++ src/tac2interp.mli | 28 + src/tac2print.ml | 296 ++++++++++ src/tac2print.mli | 37 ++ tac2core.ml | 646 --------------------- tac2core.mli | 62 -- tac2entries.ml | 645 --------------------- tac2entries.mli | 57 -- tac2env.ml | 242 -------- tac2env.mli | 106 ---- tac2expr.mli | 195 ------- tac2intern.ml | 1452 ---------------------------------------------- tac2intern.mli | 41 -- tac2interp.ml | 160 ------ tac2interp.mli | 28 - tac2print.ml | 296 ---------- tac2print.mli | 37 -- theories/Array.v | 14 + theories/Constr.v | 43 ++ theories/Control.v | 49 ++ theories/Init.v | 56 ++ theories/Int.v | 18 + theories/Ltac2.v | 16 + theories/Message.v | 20 + theories/String.v | 14 + 51 files changed, 5001 insertions(+), 4482 deletions(-) create mode 100644 .gitignore delete mode 100644 Array.v delete mode 100644 Constr.v delete mode 100644 Control.v delete mode 100644 Init.v delete mode 100644 Int.v create mode 100644 LICENSE delete mode 100644 Ltac2.v create mode 100644 Makefile delete mode 100644 Message.v create mode 100644 README.md delete mode 100644 String.v create mode 100644 _CoqProject delete mode 100644 g_ltac2.ml4 delete mode 100644 ltac2_plugin.mlpack create mode 100644 src/g_ltac2.ml4 create mode 100644 src/ltac2_plugin.mlpack create mode 100644 src/tac2core.ml create mode 100644 src/tac2core.mli create mode 100644 src/tac2entries.ml create mode 100644 src/tac2entries.mli create mode 100644 src/tac2env.ml create mode 100644 src/tac2env.mli create mode 100644 src/tac2expr.mli create mode 100644 src/tac2intern.ml create mode 100644 src/tac2intern.mli create mode 100644 src/tac2interp.ml create mode 100644 src/tac2interp.mli create mode 100644 src/tac2print.ml create mode 100644 src/tac2print.mli delete mode 100644 tac2core.ml delete mode 100644 tac2core.mli delete mode 100644 tac2entries.ml delete mode 100644 tac2entries.mli delete mode 100644 tac2env.ml delete mode 100644 tac2env.mli delete mode 100644 tac2expr.mli delete mode 100644 tac2intern.ml delete mode 100644 tac2intern.mli delete mode 100644 tac2interp.ml delete mode 100644 tac2interp.mli delete mode 100644 tac2print.ml delete mode 100644 tac2print.mli create mode 100644 theories/Array.v create mode 100644 theories/Constr.v create mode 100644 theories/Control.v create mode 100644 theories/Init.v create mode 100644 theories/Int.v create mode 100644 theories/Ltac2.v create mode 100644 theories/Message.v create mode 100644 theories/String.v diff --git a/.gitignore b/.gitignore new file mode 100644 index 0000000000..ffdea1320c --- /dev/null +++ b/.gitignore @@ -0,0 +1,14 @@ +Makefile.coq +Makefile.coq.conf +*.glob +*.d +*.d.raw +*.vio +*.vo +*.cm* +*.annot +*.spit +*.spot +*.o +*.a +*.aux diff --git a/Array.v b/Array.v deleted file mode 100644 index be4ab025ae..0000000000 --- a/Array.v +++ /dev/null @@ -1,14 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* 'a -> 'a array := "ltac2" "array_make". -Ltac2 @external length : 'a array -> int := "ltac2" "array_length". -Ltac2 @external get : 'a array -> int -> 'a := "ltac2" "array_get". -Ltac2 @external set : 'a array -> int -> 'a -> unit := "ltac2" "array_set". diff --git a/Constr.v b/Constr.v deleted file mode 100644 index 9e5833a778..0000000000 --- a/Constr.v +++ /dev/null @@ -1,43 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* constr := "ltac2" "constr_type". -(** Return the type of a term *) - -Ltac2 @ external equal : constr -> constr -> bool := "ltac2" "constr_equal". -(** Strict syntactic equality: only up to α-conversion and evar expansion *) - -Module Unsafe. - -(** Low-level access to kernel term. Use with care! *) - -Ltac2 Type kind := [ -| Rel (int) -| Var (ident) -| Meta (meta) -| Evar (evar, constr list) -| Sort (sort) -| Cast (constr, cast, constr) -| Prod (ident option, constr, constr) -| Lambda (ident option, constr, constr) -| LetIn (ident option, constr, constr, constr) -| App (constr, constr list) -| Constant (constant, instance) -| Ind (inductive, instance) -| Constructor (inductive, instance) -(* - | Case of case_info * 'constr * 'constr * 'constr array - | Fix of ('constr, 'types) pfixpoint - | CoFix of ('constr, 'types) pcofixpoint -*) -| Proj (projection, constr) -]. - -End Unsafe. diff --git a/Control.v b/Control.v deleted file mode 100644 index 6b3b360abb..0000000000 --- a/Control.v +++ /dev/null @@ -1,49 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* 'a := "ltac2" "throw". -(** Fatal exception throwing. This does not induce backtracking. *) - -(** Generic backtracking control *) - -Ltac2 @ external zero : exn -> 'a := "ltac2" "zero". -Ltac2 @ external plus : (unit -> 'a) -> (exn -> 'a) -> 'a := "ltac2" "plus". -Ltac2 @ external once : (unit -> 'a) -> 'a := "ltac2" "once". -Ltac2 @ external dispatch : (unit -> unit) list -> unit := "ltac2" "dispatch". -Ltac2 @ external extend : (unit -> unit) list -> (unit -> unit) -> (unit -> unit) list -> unit := "ltac2" "extend". -Ltac2 @ external enter : (unit -> unit) -> unit := "ltac2" "enter". - -(** Proof state manipulation *) - -Ltac2 @ external focus : int -> int -> (unit -> 'a) -> 'a := "ltac2" "focus". -Ltac2 @ external shelve : unit -> unit := "ltac2" "shelve". -Ltac2 @ external shelve_unifiable : unit -> unit := "ltac2" "shelve_unifiable". - -Ltac2 @ external new_goal : evar -> unit := "ltac2" "new_goal". -(** Adds the given evar to the list of goals as the last one. If it is - already defined in the current state, don't do anything. Panics if the - evar is not in the current state. *) - -(** Goal inspection *) - -Ltac2 @ external goal : unit -> constr := "ltac2" "goal". -(** Panics if there is not exactly one goal under focus. Otherwise returns - the conclusion of this goal. *) - -Ltac2 @ external hyp : ident -> constr := "ltac2" "hyp". -(** Panics if there is more than one goal under focus. If there is no - goal under focus, looks for the section variable with the given name. - If there is one, looks for the hypothesis with the given name. *) - -(** Refinement *) - -Ltac2 @ external refine : (unit -> constr) -> unit := "ltac2" "refine". diff --git a/Init.v b/Init.v deleted file mode 100644 index 1d2d40f5c0..0000000000 --- a/Init.v +++ /dev/null @@ -1,56 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* int -> bool := "ltac2" "int_equal". -Ltac2 @ external compare : int -> int -> int := "ltac2" "int_compare". -Ltac2 @ external add : int -> int -> int := "ltac2" "int_add". -Ltac2 @ external sub : int -> int -> int := "ltac2" "int_sub". -Ltac2 @ external mul : int -> int -> int := "ltac2" "int_mul". -Ltac2 @ external neg : int -> int := "ltac2" "int_neg". diff --git a/LICENSE b/LICENSE new file mode 100644 index 0000000000..27950e8d20 --- /dev/null +++ b/LICENSE @@ -0,0 +1,458 @@ + GNU LESSER GENERAL PUBLIC LICENSE + Version 2.1, February 1999 + + Copyright (C) 1991, 1999 Free Software Foundation, Inc. + 59 Temple Place, Suite 330, Boston, MA 02111-1307 USA + Everyone is permitted to copy and distribute verbatim copies + of this license document, but changing it is not allowed. + +[This is the first released version of the Lesser GPL. It also counts + as the successor of the GNU Library Public License, version 2, hence + the version number 2.1.] + + Preamble + + The licenses for most software are designed to take away your +freedom to share and change it. By contrast, the GNU General Public +Licenses are intended to guarantee your freedom to share and change +free software--to make sure the software is free for all its users. + + This license, the Lesser General Public License, applies to some +specially designated software packages--typically libraries--of the +Free Software Foundation and other authors who decide to use it. You +can use it too, but we suggest you first think carefully about whether +this license or the ordinary General Public License is the better +strategy to use in any particular case, based on the explanations below. + + When we speak of free software, we are referring to freedom of use, +not price. Our General Public Licenses are designed to make sure that +you have the freedom to distribute copies of free software (and charge +for this service if you wish); that you receive source code or can get +it if you want it; that you can change the software and use pieces of +it in new free programs; and that you are informed that you can do +these things. + + To protect your rights, we need to make restrictions that forbid +distributors to deny you these rights or to ask you to surrender these +rights. These restrictions translate to certain responsibilities for +you if you distribute copies of the library or if you modify it. + + For example, if you distribute copies of the library, whether gratis +or for a fee, you must give the recipients all the rights that we gave +you. You must make sure that they, too, receive or can get the source +code. If you link other code with the library, you must provide +complete object files to the recipients, so that they can relink them +with the library after making changes to the library and recompiling +it. And you must show them these terms so they know their rights. + + We protect your rights with a two-step method: (1) we copyright the +library, and (2) we offer you this license, which gives you legal +permission to copy, distribute and/or modify the library. + + To protect each distributor, we want to make it very clear that +there is no warranty for the free library. Also, if the library is +modified by someone else and passed on, the recipients should know +that what they have is not the original version, so that the original +author's reputation will not be affected by problems that might be +introduced by others. + + Finally, software patents pose a constant threat to the existence of +any free program. We wish to make sure that a company cannot +effectively restrict the users of a free program by obtaining a +restrictive license from a patent holder. Therefore, we insist that +any patent license obtained for a version of the library must be +consistent with the full freedom of use specified in this license. + + Most GNU software, including some libraries, is covered by the +ordinary GNU General Public License. This license, the GNU Lesser +General Public License, applies to certain designated libraries, and +is quite different from the ordinary General Public License. We use +this license for certain libraries in order to permit linking those +libraries into non-free programs. + + When a program is linked with a library, whether statically or using +a shared library, the combination of the two is legally speaking a +combined work, a derivative of the original library. The ordinary +General Public License therefore permits such linking only if the +entire combination fits its criteria of freedom. The Lesser General +Public License permits more lax criteria for linking other code with +the library. + + We call this license the "Lesser" General Public License because it +does Less to protect the user's freedom than the ordinary General +Public License. It also provides other free software developers Less +of an advantage over competing non-free programs. These disadvantages +are the reason we use the ordinary General Public License for many +libraries. However, the Lesser license provides advantages in certain +special circumstances. + + For example, on rare occasions, there may be a special need to +encourage the widest possible use of a certain library, so that it becomes +a de-facto standard. To achieve this, non-free programs must be +allowed to use the library. A more frequent case is that a free +library does the same job as widely used non-free libraries. In this +case, there is little to gain by limiting the free library to free +software only, so we use the Lesser General Public License. + + In other cases, permission to use a particular library in non-free +programs enables a greater number of people to use a large body of +free software. For example, permission to use the GNU C Library in +non-free programs enables many more people to use the whole GNU +operating system, as well as its variant, the GNU/Linux operating +system. + + Although the Lesser General Public License is Less protective of the +users' freedom, it does ensure that the user of a program that is +linked with the Library has the freedom and the wherewithal to run +that program using a modified version of the Library. + + The precise terms and conditions for copying, distribution and +modification follow. Pay close attention to the difference between a +"work based on the library" and a "work that uses the library". The +former contains code derived from the library, whereas the latter must +be combined with the library in order to run. + + GNU LESSER GENERAL PUBLIC LICENSE + TERMS AND CONDITIONS FOR COPYING, DISTRIBUTION AND MODIFICATION + + 0. This License Agreement applies to any software library or other +program which contains a notice placed by the copyright holder or +other authorized party saying it may be distributed under the terms of +this Lesser General Public License (also called "this License"). +Each licensee is addressed as "you". + + A "library" means a collection of software functions and/or data +prepared so as to be conveniently linked with application programs +(which use some of those functions and data) to form executables. + + The "Library", below, refers to any such software library or work +which has been distributed under these terms. A "work based on the +Library" means either the Library or any derivative work under +copyright law: that is to say, a work containing the Library or a +portion of it, either verbatim or with modifications and/or translated +straightforwardly into another language. (Hereinafter, translation is +included without limitation in the term "modification".) + + "Source code" for a work means the preferred form of the work for +making modifications to it. For a library, complete source code means +all the source code for all modules it contains, plus any associated +interface definition files, plus the scripts used to control compilation +and installation of the library. + + Activities other than copying, distribution and modification are not +covered by this License; they are outside its scope. The act of +running a program using the Library is not restricted, and output from +such a program is covered only if its contents constitute a work based +on the Library (independent of the use of the Library in a tool for +writing it). Whether that is true depends on what the Library does +and what the program that uses the Library does. + + 1. You may copy and distribute verbatim copies of the Library's +complete source code as you receive it, in any medium, provided that +you conspicuously and appropriately publish on each copy an +appropriate copyright notice and disclaimer of warranty; keep intact +all the notices that refer to this License and to the absence of any +warranty; and distribute a copy of this License along with the +Library. + + You may charge a fee for the physical act of transferring a copy, +and you may at your option offer warranty protection in exchange for a +fee. + + 2. You may modify your copy or copies of the Library or any portion +of it, thus forming a work based on the Library, and copy and +distribute such modifications or work under the terms of Section 1 +above, provided that you also meet all of these conditions: + + a) The modified work must itself be a software library. + + b) You must cause the files modified to carry prominent notices + stating that you changed the files and the date of any change. + + c) You must cause the whole of the work to be licensed at no + charge to all third parties under the terms of this License. + + d) If a facility in the modified Library refers to a function or a + table of data to be supplied by an application program that uses + the facility, other than as an argument passed when the facility + is invoked, then you must make a good faith effort to ensure that, + in the event an application does not supply such function or + table, the facility still operates, and performs whatever part of + its purpose remains meaningful. + + (For example, a function in a library to compute square roots has + a purpose that is entirely well-defined independent of the + application. Therefore, Subsection 2d requires that any + application-supplied function or table used by this function must + be optional: if the application does not supply it, the square + root function must still compute square roots.) + +These requirements apply to the modified work as a whole. If +identifiable sections of that work are not derived from the Library, +and can be reasonably considered independent and separate works in +themselves, then this License, and its terms, do not apply to those +sections when you distribute them as separate works. But when you +distribute the same sections as part of a whole which is a work based +on the Library, the distribution of the whole must be on the terms of +this License, whose permissions for other licensees extend to the +entire whole, and thus to each and every part regardless of who wrote +it. + +Thus, it is not the intent of this section to claim rights or contest +your rights to work written entirely by you; rather, the intent is to +exercise the right to control the distribution of derivative or +collective works based on the Library. + +In addition, mere aggregation of another work not based on the Library +with the Library (or with a work based on the Library) on a volume of +a storage or distribution medium does not bring the other work under +the scope of this License. + + 3. You may opt to apply the terms of the ordinary GNU General Public +License instead of this License to a given copy of the Library. To do +this, you must alter all the notices that refer to this License, so +that they refer to the ordinary GNU General Public License, version 2, +instead of to this License. (If a newer version than version 2 of the +ordinary GNU General Public License has appeared, then you can specify +that version instead if you wish.) Do not make any other change in +these notices. + + Once this change is made in a given copy, it is irreversible for +that copy, so the ordinary GNU General Public License applies to all +subsequent copies and derivative works made from that copy. + + This option is useful when you wish to copy part of the code of +the Library into a program that is not a library. + + 4. You may copy and distribute the Library (or a portion or +derivative of it, under Section 2) in object code or executable form +under the terms of Sections 1 and 2 above provided that you accompany +it with the complete corresponding machine-readable source code, which +must be distributed under the terms of Sections 1 and 2 above on a +medium customarily used for software interchange. + + If distribution of object code is made by offering access to copy +from a designated place, then offering equivalent access to copy the +source code from the same place satisfies the requirement to +distribute the source code, even though third parties are not +compelled to copy the source along with the object code. + + 5. A program that contains no derivative of any portion of the +Library, but is designed to work with the Library by being compiled or +linked with it, is called a "work that uses the Library". Such a +work, in isolation, is not a derivative work of the Library, and +therefore falls outside the scope of this License. + + However, linking a "work that uses the Library" with the Library +creates an executable that is a derivative of the Library (because it +contains portions of the Library), rather than a "work that uses the +library". The executable is therefore covered by this License. +Section 6 states terms for distribution of such executables. + + When a "work that uses the Library" uses material from a header file +that is part of the Library, the object code for the work may be a +derivative work of the Library even though the source code is not. +Whether this is true is especially significant if the work can be +linked without the Library, or if the work is itself a library. The +threshold for this to be true is not precisely defined by law. + + If such an object file uses only numerical parameters, data +structure layouts and accessors, and small macros and small inline +functions (ten lines or less in length), then the use of the object +file is unrestricted, regardless of whether it is legally a derivative +work. (Executables containing this object code plus portions of the +Library will still fall under Section 6.) + + Otherwise, if the work is a derivative of the Library, you may +distribute the object code for the work under the terms of Section 6. +Any executables containing that work also fall under Section 6, +whether or not they are linked directly with the Library itself. + + 6. As an exception to the Sections above, you may also combine or +link a "work that uses the Library" with the Library to produce a +work containing portions of the Library, and distribute that work +under terms of your choice, provided that the terms permit +modification of the work for the customer's own use and reverse +engineering for debugging such modifications. + + You must give prominent notice with each copy of the work that the +Library is used in it and that the Library and its use are covered by +this License. You must supply a copy of this License. If the work +during execution displays copyright notices, you must include the +copyright notice for the Library among them, as well as a reference +directing the user to the copy of this License. Also, you must do one +of these things: + + a) Accompany the work with the complete corresponding + machine-readable source code for the Library including whatever + changes were used in the work (which must be distributed under + Sections 1 and 2 above); and, if the work is an executable linked + with the Library, with the complete machine-readable "work that + uses the Library", as object code and/or source code, so that the + user can modify the Library and then relink to produce a modified + executable containing the modified Library. (It is understood + that the user who changes the contents of definitions files in the + Library will not necessarily be able to recompile the application + to use the modified definitions.) + + b) Use a suitable shared library mechanism for linking with the + Library. A suitable mechanism is one that (1) uses at run time a + copy of the library already present on the user's computer system, + rather than copying library functions into the executable, and (2) + will operate properly with a modified version of the library, if + the user installs one, as long as the modified version is + interface-compatible with the version that the work was made with. + + c) Accompany the work with a written offer, valid for at + least three years, to give the same user the materials + specified in Subsection 6a, above, for a charge no more + than the cost of performing this distribution. + + d) If distribution of the work is made by offering access to copy + from a designated place, offer equivalent access to copy the above + specified materials from the same place. + + e) Verify that the user has already received a copy of these + materials or that you have already sent this user a copy. + + For an executable, the required form of the "work that uses the +Library" must include any data and utility programs needed for +reproducing the executable from it. However, as a special exception, +the materials to be distributed need not include anything that is +normally distributed (in either source or binary form) with the major +components (compiler, kernel, and so on) of the operating system on +which the executable runs, unless that component itself accompanies +the executable. + + It may happen that this requirement contradicts the license +restrictions of other proprietary libraries that do not normally +accompany the operating system. Such a contradiction means you cannot +use both them and the Library together in an executable that you +distribute. + + 7. You may place library facilities that are a work based on the +Library side-by-side in a single library together with other library +facilities not covered by this License, and distribute such a combined +library, provided that the separate distribution of the work based on +the Library and of the other library facilities is otherwise +permitted, and provided that you do these two things: + + a) Accompany the combined library with a copy of the same work + based on the Library, uncombined with any other library + facilities. This must be distributed under the terms of the + Sections above. + + b) Give prominent notice with the combined library of the fact + that part of it is a work based on the Library, and explaining + where to find the accompanying uncombined form of the same work. + + 8. You may not copy, modify, sublicense, link with, or distribute +the Library except as expressly provided under this License. Any +attempt otherwise to copy, modify, sublicense, link with, or +distribute the Library is void, and will automatically terminate your +rights under this License. However, parties who have received copies, +or rights, from you under this License will not have their licenses +terminated so long as such parties remain in full compliance. + + 9. You are not required to accept this License, since you have not +signed it. However, nothing else grants you permission to modify or +distribute the Library or its derivative works. These actions are +prohibited by law if you do not accept this License. Therefore, by +modifying or distributing the Library (or any work based on the +Library), you indicate your acceptance of this License to do so, and +all its terms and conditions for copying, distributing or modifying +the Library or works based on it. + + 10. Each time you redistribute the Library (or any work based on the +Library), the recipient automatically receives a license from the +original licensor to copy, distribute, link with or modify the Library +subject to these terms and conditions. You may not impose any further +restrictions on the recipients' exercise of the rights granted herein. +You are not responsible for enforcing compliance by third parties with +this License. + + 11. If, as a consequence of a court judgment or allegation of patent +infringement or for any other reason (not limited to patent issues), +conditions are imposed on you (whether by court order, agreement or +otherwise) that contradict the conditions of this License, they do not +excuse you from the conditions of this License. If you cannot +distribute so as to satisfy simultaneously your obligations under this +License and any other pertinent obligations, then as a consequence you +may not distribute the Library at all. For example, if a patent +license would not permit royalty-free redistribution of the Library by +all those who receive copies directly or indirectly through you, then +the only way you could satisfy both it and this License would be to +refrain entirely from distribution of the Library. + +If any portion of this section is held invalid or unenforceable under any +particular circumstance, the balance of the section is intended to apply, +and the section as a whole is intended to apply in other circumstances. + +It is not the purpose of this section to induce you to infringe any +patents or other property right claims or to contest validity of any +such claims; this section has the sole purpose of protecting the +integrity of the free software distribution system which is +implemented by public license practices. Many people have made +generous contributions to the wide range of software distributed +through that system in reliance on consistent application of that +system; it is up to the author/donor to decide if he or she is willing +to distribute software through any other system and a licensee cannot +impose that choice. + +This section is intended to make thoroughly clear what is believed to +be a consequence of the rest of this License. + + 12. If the distribution and/or use of the Library is restricted in +certain countries either by patents or by copyrighted interfaces, the +original copyright holder who places the Library under this License may add +an explicit geographical distribution limitation excluding those countries, +so that distribution is permitted only in or among countries not thus +excluded. In such case, this License incorporates the limitation as if +written in the body of this License. + + 13. The Free Software Foundation may publish revised and/or new +versions of the Lesser General Public License from time to time. +Such new versions will be similar in spirit to the present version, +but may differ in detail to address new problems or concerns. + +Each version is given a distinguishing version number. If the Library +specifies a version number of this License which applies to it and +"any later version", you have the option of following the terms and +conditions either of that version or of any later version published by +the Free Software Foundation. If the Library does not specify a +license version number, you may choose any version ever published by +the Free Software Foundation. + + 14. If you wish to incorporate parts of the Library into other free +programs whose distribution conditions are incompatible with these, +write to the author to ask for permission. For software which is +copyrighted by the Free Software Foundation, write to the Free +Software Foundation; we sometimes make exceptions for this. Our +decision will be guided by the two goals of preserving the free status +of all derivatives of our free software and of promoting the sharing +and reuse of software generally. + + NO WARRANTY + + 15. BECAUSE THE LIBRARY IS LICENSED FREE OF CHARGE, THERE IS NO +WARRANTY FOR THE LIBRARY, TO THE EXTENT PERMITTED BY APPLICABLE LAW. +EXCEPT WHEN OTHERWISE STATED IN WRITING THE COPYRIGHT HOLDERS AND/OR +OTHER PARTIES PROVIDE THE LIBRARY "AS IS" WITHOUT WARRANTY OF ANY +KIND, EITHER EXPRESSED OR IMPLIED, INCLUDING, BUT NOT LIMITED TO, THE +IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR +PURPOSE. THE ENTIRE RISK AS TO THE QUALITY AND PERFORMANCE OF THE +LIBRARY IS WITH YOU. SHOULD THE LIBRARY PROVE DEFECTIVE, YOU ASSUME +THE COST OF ALL NECESSARY SERVICING, REPAIR OR CORRECTION. + + 16. IN NO EVENT UNLESS REQUIRED BY APPLICABLE LAW OR AGREED TO IN +WRITING WILL ANY COPYRIGHT HOLDER, OR ANY OTHER PARTY WHO MAY MODIFY +AND/OR REDISTRIBUTE THE LIBRARY AS PERMITTED ABOVE, BE LIABLE TO YOU +FOR DAMAGES, INCLUDING ANY GENERAL, SPECIAL, INCIDENTAL OR +CONSEQUENTIAL DAMAGES ARISING OUT OF THE USE OR INABILITY TO USE THE +LIBRARY (INCLUDING BUT NOT LIMITED TO LOSS OF DATA OR DATA BEING +RENDERED INACCURATE OR LOSSES SUSTAINED BY YOU OR THIRD PARTIES OR A +FAILURE OF THE LIBRARY TO OPERATE WITH ANY OTHER SOFTWARE), EVEN IF +SUCH HOLDER OR OTHER PARTY HAS BEEN ADVISED OF THE POSSIBILITY OF SUCH +DAMAGES. + + END OF TERMS AND CONDITIONS diff --git a/Ltac2.v b/Ltac2.v deleted file mode 100644 index 625d4ac513..0000000000 --- a/Ltac2.v +++ /dev/null @@ -1,16 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* unit := "ltac2" "print". - -Ltac2 @ external of_string : string -> message := "ltac2" "message_of_string". - -Ltac2 @ external of_int : int -> message := "ltac2" "message_of_int". - -Ltac2 @ external of_constr : constr -> message := "ltac2" "message_of_constr". -(** Panics if there is more than one goal under focus. *) - -Ltac2 @ external concat : message -> message -> message := "ltac2" "message_concat". diff --git a/README.md b/README.md new file mode 100644 index 0000000000..e69de29bb2 diff --git a/String.v b/String.v deleted file mode 100644 index 3a4a178878..0000000000 --- a/String.v +++ /dev/null @@ -1,14 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* char -> string := "ltac2" "string_make". -Ltac2 @external length : string -> int := "ltac2" "string_length". -Ltac2 @external get : string -> int -> char := "ltac2" "string_get". -Ltac2 @external set : string -> int -> char -> unit := "ltac2" "string_set". diff --git a/_CoqProject b/_CoqProject new file mode 100644 index 0000000000..90338abbfb --- /dev/null +++ b/_CoqProject @@ -0,0 +1,28 @@ +-R theories/ Ltac2 +-I src/ +-bypass-API + +src/tac2expr.mli +src/tac2env.ml +src/tac2env.mli +src/tac2print.ml +src/tac2print.mli +src/tac2intern.ml +src/tac2intern.mli +src/tac2interp.ml +src/tac2interp.mli +src/tac2entries.ml +src/tac2entries.mli +src/tac2core.ml +src/tac2core.mli +src/g_ltac2.ml4 +src/ltac2_plugin.mlpack + +theories/Init.v +theories/Int.v +theories/String.v +theories/Array.v +theories/Control.v +theories/Message.v +theories/Constr.v +theories/Ltac2.v diff --git a/g_ltac2.ml4 b/g_ltac2.ml4 deleted file mode 100644 index 51919addf2..0000000000 --- a/g_ltac2.ml4 +++ /dev/null @@ -1,278 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* CPatRef (!@loc, RelId id, pl) - | id = Prim.qualid -> CPatRef (!@loc, RelId id, []) ] - | "0" - [ "_" -> CPatAny (!@loc) - | "()" -> CPatTup (!@loc, []) - | id = Prim.qualid -> CPatRef (!@loc, RelId id, []) - | "("; pl = LIST0 tac2pat LEVEL "1" SEP ","; ")" -> CPatTup (!@loc, pl) ] - ] - ; - tac2expr: - [ "5" - [ "fun"; it = LIST1 input_fun ; "=>"; body = tac2expr LEVEL "5" -> CTacFun (!@loc, it, body) - | "let"; isrec = rec_flag; - lc = LIST1 let_clause SEP "with"; "in"; - e = tac2expr LEVEL "5" -> CTacLet (!@loc, isrec, lc, e) - | "match"; e = tac2expr LEVEL "5"; "with"; bl = branches ;"end" -> - CTacCse (!@loc, e, bl) - ] - | "2" LEFTA - [ e1 = tac2expr; ";"; e2 = tac2expr -> CTacSeq (!@loc, e1, e2) ] - | "1" LEFTA - [ e = tac2expr; el = LIST1 tac2expr LEVEL "0" -> CTacApp (!@loc, e, el) - | e = SELF; ".("; qid = Prim.qualid; ")" -> CTacPrj (!@loc, e, RelId qid) - | e = SELF; ".("; qid = Prim.qualid; ")"; ":="; r = tac2expr LEVEL "1" -> CTacSet (!@loc, e, RelId qid, r) - | e0 = tac2expr; ","; el = LIST1 tac2expr LEVEL "1" SEP "," -> CTacTup (!@loc, e0 :: el) ] - | "0" - [ "("; a = tac2expr LEVEL "5"; ")" -> a - | "("; a = tac2expr; ":"; t = tac2type; ")" -> CTacCnv (!@loc, a, t) - | "()" -> CTacTup (!@loc, []) - | "("; ")" -> CTacTup (!@loc, []) - | "["; a = LIST0 tac2expr LEVEL "1" SEP ";"; "]" -> CTacLst (!@loc, a) - | "{"; a = tac2rec_fieldexprs; "}" -> CTacRec (!@loc, a) - | a = tactic_atom -> a ] - ] - ; - branches: - [ [ -> [] - | "|"; bl = LIST1 branch SEP "|" -> bl - | bl = LIST1 branch SEP "|" -> bl ] - ] - ; - branch: - [ [ pat = tac2pat LEVEL "1"; "=>"; e = tac2expr LEVEL "5" -> (pat, e) ] ] - ; - rec_flag: - [ [ IDENT "rec" -> true - | -> false ] ] - ; - typ_param: - [ [ "'"; id = Prim.ident -> id ] ] - ; - tactic_atom: - [ [ n = Prim.integer -> CTacAtm (!@loc, AtmInt n) - | s = Prim.string -> CTacAtm (!@loc, AtmStr s) - | id = Prim.qualid -> CTacRef (RelId id) - | IDENT "constr"; ":"; "("; c = Constr.lconstr; ")" -> inj_constr !@loc c - | IDENT "open_constr"; ":"; "("; c = Constr.lconstr; ")" -> inj_open_constr !@loc c - | IDENT "ident"; ":"; "("; c = Prim.ident; ")" -> inj_ident !@loc c - ] ] - ; - let_clause: - [ [ id = binder; ":="; te = tac2expr -> - (id, None, te) - | id = binder; args = LIST1 input_fun; ":="; te = tac2expr -> - (id, None, CTacFun (!@loc, args, te)) ] ] - ; - tac2type: - [ "5" RIGHTA - [ t1 = tac2type; "->"; t2 = tac2type -> CTypArrow (!@loc, t1, t2) ] - | "2" - [ t = tac2type; "*"; tl = LIST1 tac2type SEP "*" -> CTypTuple (!@loc, t :: tl) ] - | "1" LEFTA - [ t = SELF; qid = Prim.qualid -> CTypRef (!@loc, RelId qid, [t]) ] - | "0" - [ "("; t = tac2type LEVEL "5"; ")" -> t - | id = typ_param -> CTypVar (!@loc, Name id) - | "_" -> CTypVar (!@loc, Anonymous) - | qid = Prim.qualid -> CTypRef (!@loc, RelId qid, []) - | "("; p = LIST1 tac2type LEVEL "5" SEP ","; ")"; qid = Prim.qualid -> - CTypRef (!@loc, RelId qid, p) ] - ]; - locident: - [ [ id = Prim.ident -> (!@loc, id) ] ] - ; - binder: - [ [ "_" -> (!@loc, Anonymous) - | l = Prim.ident -> (!@loc, Name l) ] ] - ; - input_fun: - [ [ b = binder -> (b, None) - | "("; b = binder; ":"; t = tac2type; ")" -> (b, Some t) ] ] - ; - tac2def_body: - [ [ name = binder; it = LIST0 input_fun; ":="; e = tac2expr -> - let e = if List.is_empty it then e else CTacFun (!@loc, it, e) in - (name, e) - ] ] - ; - tac2def_val: - [ [ isrec = rec_flag; l = LIST1 tac2def_body SEP "with" -> - StrVal (isrec, l) - ] ] - ; - tac2typ_knd: - [ [ t = tac2type -> CTydDef (Some t) - | "["; ".."; "]" -> CTydOpn - | "["; t = tac2alg_constructors; "]" -> CTydAlg t - | "{"; t = tac2rec_fields; "}"-> CTydRec t ] ] - ; - tac2alg_constructors: - [ [ "|"; cs = LIST1 tac2alg_constructor SEP "|" -> cs - | cs = LIST0 tac2alg_constructor SEP "|" -> cs ] ] - ; - tac2alg_constructor: - [ [ c = Prim.ident -> (c, []) - | c = Prim.ident; "("; args = LIST0 tac2type SEP ","; ")"-> (c, args) ] ] - ; - tac2rec_fields: - [ [ f = tac2rec_field; ";"; l = tac2rec_fields -> f :: l - | f = tac2rec_field; ";" -> [f] - | f = tac2rec_field -> [f] - | -> [] ] ] - ; - tac2rec_field: - [ [ mut = [ -> false | IDENT "mutable" -> true]; id = Prim.ident; ":"; t = tac2type -> (id, mut, t) ] ] - ; - tac2rec_fieldexprs: - [ [ f = tac2rec_fieldexpr; ";"; l = tac2rec_fieldexprs -> f :: l - | f = tac2rec_fieldexpr; ";" -> [f] - | f = tac2rec_fieldexpr-> [f] - | -> [] ] ] - ; - tac2rec_fieldexpr: - [ [ qid = Prim.qualid; ":="; e = tac2expr LEVEL "1" -> RelId qid, e ] ] - ; - tac2typ_prm: - [ [ -> [] - | id = typ_param -> [!@loc, id] - | "("; ids = LIST1 [ id = typ_param -> (!@loc, id) ] SEP "," ;")" -> ids - ] ] - ; - tac2typ_def: - [ [ prm = tac2typ_prm; id = Prim.qualid; (r, e) = tac2type_body -> (id, r, (prm, e)) ] ] - ; - tac2type_body: - [ [ -> false, CTydDef None - | ":="; e = tac2typ_knd -> false, e - | "::="; e = tac2typ_knd -> true, e - ] ] - ; - tac2def_typ: - [ [ "Type"; isrec = rec_flag; l = LIST1 tac2typ_def SEP "with" -> - StrTyp (isrec, l) - ] ] - ; - tac2def_ext: - [ [ "@"; IDENT "external"; id = locident; ":"; t = tac2type LEVEL "5"; ":="; - plugin = Prim.string; name = Prim.string -> - let ml = { mltac_plugin = plugin; mltac_tactic = name } in - StrPrm (id, t, ml) - ] ] - ; - syn_node: - [ [ "_" -> (!@loc, None) - | id = Prim.ident -> (!@loc, Some id) - ] ] - ; - sexpr: - [ [ s = Prim.string -> SexprStr (!@loc, s) - | n = Prim.integer -> SexprInt (!@loc, n) - | id = syn_node -> SexprRec (!@loc, id, []) - | id = syn_node; "("; tok = LIST1 sexpr SEP "," ; ")" -> - SexprRec (!@loc, id, tok) - ] ] - ; - syn_level: - [ [ -> None - | ":"; n = Prim.integer -> Some n - ] ] - ; - tac2def_syn: - [ [ "Notation"; toks = LIST1 sexpr; n = syn_level; ":="; - e = tac2expr -> - StrSyn (toks, n, e) - ] ] - ; -END - -GEXTEND Gram - Pcoq.Constr.operconstr: LEVEL "0" - [ [ IDENT "ltac2"; ":"; "("; tac = tac2expr; ")" -> - let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in - CHole (!@loc, None, IntroAnonymous, Some arg) ] ] - ; -END - -let pr_ltac2entry _ = mt () (** FIXME *) -let pr_ltac2expr _ = mt () (** FIXME *) - -VERNAC ARGUMENT EXTEND ltac2_entry -PRINTED BY pr_ltac2entry -| [ tac2def_val(v) ] -> [ v ] -| [ tac2def_typ(t) ] -> [ t ] -| [ tac2def_ext(e) ] -> [ e ] -| [ tac2def_syn(e) ] -> [ e ] -END - -VERNAC COMMAND EXTEND VernacDeclareTactic2Definition CLASSIFIED AS SIDEFF -| [ "Ltac2" ltac2_entry(e) ] -> [ - let local = Locality.LocalityFixme.consume () in - Tac2entries.register_struct ?local e - ] -END - -let _ = - let mode = { - Proof_global.name = "Ltac2"; - set = (fun () -> set_command_entry tac2mode); - reset = (fun () -> set_command_entry Pcoq.Vernac_.noedit_mode); - } in - Proof_global.register_proof_mode mode - -VERNAC ARGUMENT EXTEND ltac2_expr -PRINTED BY pr_ltac2expr -| [ tac2expr(e) ] -> [ e ] -END - -open G_ltac -open Vernac_classifier - -VERNAC tac2mode EXTEND VernacLtac2 -| [ - ltac2_expr(t) ltac_use_default(default) ] => - [ classify_as_proofstep ] -> [ -(* let g = Option.default (Proof_global.get_default_goal_selector ()) g in *) - Tac2entries.call ~default t - ] -END - -open Stdarg - -VERNAC COMMAND EXTEND Ltac2Print CLASSIFIED AS SIDEFF -| [ "Print" "Ltac2" reference(tac) ] -> [ Tac2entries.print_ltac tac ] -END - diff --git a/ltac2_plugin.mlpack b/ltac2_plugin.mlpack deleted file mode 100644 index 3d87a8cddb..0000000000 --- a/ltac2_plugin.mlpack +++ /dev/null @@ -1,7 +0,0 @@ -Tac2env -Tac2print -Tac2intern -Tac2interp -Tac2entries -Tac2core -G_ltac2 diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 new file mode 100644 index 0000000000..36057b3a67 --- /dev/null +++ b/src/g_ltac2.ml4 @@ -0,0 +1,278 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* CPatRef (!@loc, RelId id, pl) + | id = Prim.qualid -> CPatRef (!@loc, RelId id, []) ] + | "0" + [ "_" -> CPatAny (!@loc) + | "()" -> CPatTup (Loc.tag ~loc:!@loc []) + | id = Prim.qualid -> CPatRef (!@loc, RelId id, []) + | "("; pl = LIST0 tac2pat LEVEL "1" SEP ","; ")" -> CPatTup (Loc.tag ~loc:!@loc pl) ] + ] + ; + tac2expr: + [ "5" + [ "fun"; it = LIST1 input_fun ; "=>"; body = tac2expr LEVEL "5" -> CTacFun (!@loc, it, body) + | "let"; isrec = rec_flag; + lc = LIST1 let_clause SEP "with"; "in"; + e = tac2expr LEVEL "5" -> CTacLet (!@loc, isrec, lc, e) + | "match"; e = tac2expr LEVEL "5"; "with"; bl = branches ;"end" -> + CTacCse (!@loc, e, bl) + ] + | "2" LEFTA + [ e1 = tac2expr; ";"; e2 = tac2expr -> CTacSeq (!@loc, e1, e2) ] + | "1" LEFTA + [ e = tac2expr; el = LIST1 tac2expr LEVEL "0" -> CTacApp (!@loc, e, el) + | e = SELF; ".("; qid = Prim.qualid; ")" -> CTacPrj (!@loc, e, RelId qid) + | e = SELF; ".("; qid = Prim.qualid; ")"; ":="; r = tac2expr LEVEL "1" -> CTacSet (!@loc, e, RelId qid, r) + | e0 = tac2expr; ","; el = LIST1 tac2expr LEVEL "1" SEP "," -> CTacTup (Loc.tag ~loc:!@loc (e0 :: el)) ] + | "0" + [ "("; a = tac2expr LEVEL "5"; ")" -> a + | "("; a = tac2expr; ":"; t = tac2type; ")" -> CTacCnv (!@loc, a, t) + | "()" -> CTacTup (Loc.tag ~loc:!@loc []) + | "("; ")" -> CTacTup (Loc.tag ~loc:!@loc []) + | "["; a = LIST0 tac2expr LEVEL "1" SEP ";"; "]" -> CTacLst (Loc.tag ~loc:!@loc a) + | "{"; a = tac2rec_fieldexprs; "}" -> CTacRec (!@loc, a) + | a = tactic_atom -> a ] + ] + ; + branches: + [ [ -> [] + | "|"; bl = LIST1 branch SEP "|" -> bl + | bl = LIST1 branch SEP "|" -> bl ] + ] + ; + branch: + [ [ pat = tac2pat LEVEL "1"; "=>"; e = tac2expr LEVEL "5" -> (pat, e) ] ] + ; + rec_flag: + [ [ IDENT "rec" -> true + | -> false ] ] + ; + typ_param: + [ [ "'"; id = Prim.ident -> id ] ] + ; + tactic_atom: + [ [ n = Prim.integer -> CTacAtm (Loc.tag ~loc:!@loc (AtmInt n)) + | s = Prim.string -> CTacAtm (Loc.tag ~loc:!@loc (AtmStr s)) + | id = Prim.qualid -> CTacRef (RelId id) + | IDENT "constr"; ":"; "("; c = Constr.lconstr; ")" -> inj_constr !@loc c + | IDENT "open_constr"; ":"; "("; c = Constr.lconstr; ")" -> inj_open_constr !@loc c + | IDENT "ident"; ":"; "("; c = Prim.ident; ")" -> inj_ident !@loc c + ] ] + ; + let_clause: + [ [ id = binder; ":="; te = tac2expr -> + (id, None, te) + | id = binder; args = LIST1 input_fun; ":="; te = tac2expr -> + (id, None, CTacFun (!@loc, args, te)) ] ] + ; + tac2type: + [ "5" RIGHTA + [ t1 = tac2type; "->"; t2 = tac2type -> CTypArrow (!@loc, t1, t2) ] + | "2" + [ t = tac2type; "*"; tl = LIST1 tac2type SEP "*" -> CTypTuple (!@loc, t :: tl) ] + | "1" LEFTA + [ t = SELF; qid = Prim.qualid -> CTypRef (!@loc, RelId qid, [t]) ] + | "0" + [ "("; t = tac2type LEVEL "5"; ")" -> t + | id = typ_param -> CTypVar (Loc.tag ~loc:!@loc (Name id)) + | "_" -> CTypVar (Loc.tag ~loc:!@loc Anonymous) + | qid = Prim.qualid -> CTypRef (!@loc, RelId qid, []) + | "("; p = LIST1 tac2type LEVEL "5" SEP ","; ")"; qid = Prim.qualid -> + CTypRef (!@loc, RelId qid, p) ] + ]; + locident: + [ [ id = Prim.ident -> Loc.tag ~loc:!@loc id ] ] + ; + binder: + [ [ "_" -> Loc.tag ~loc:!@loc Anonymous + | l = Prim.ident -> Loc.tag ~loc:!@loc (Name l) ] ] + ; + input_fun: + [ [ b = binder -> (b, None) + | "("; b = binder; ":"; t = tac2type; ")" -> (b, Some t) ] ] + ; + tac2def_body: + [ [ name = binder; it = LIST0 input_fun; ":="; e = tac2expr -> + let e = if List.is_empty it then e else CTacFun (!@loc, it, e) in + (name, e) + ] ] + ; + tac2def_val: + [ [ isrec = rec_flag; l = LIST1 tac2def_body SEP "with" -> + StrVal (isrec, l) + ] ] + ; + tac2typ_knd: + [ [ t = tac2type -> CTydDef (Some t) + | "["; ".."; "]" -> CTydOpn + | "["; t = tac2alg_constructors; "]" -> CTydAlg t + | "{"; t = tac2rec_fields; "}"-> CTydRec t ] ] + ; + tac2alg_constructors: + [ [ "|"; cs = LIST1 tac2alg_constructor SEP "|" -> cs + | cs = LIST0 tac2alg_constructor SEP "|" -> cs ] ] + ; + tac2alg_constructor: + [ [ c = Prim.ident -> (c, []) + | c = Prim.ident; "("; args = LIST0 tac2type SEP ","; ")"-> (c, args) ] ] + ; + tac2rec_fields: + [ [ f = tac2rec_field; ";"; l = tac2rec_fields -> f :: l + | f = tac2rec_field; ";" -> [f] + | f = tac2rec_field -> [f] + | -> [] ] ] + ; + tac2rec_field: + [ [ mut = [ -> false | IDENT "mutable" -> true]; id = Prim.ident; ":"; t = tac2type -> (id, mut, t) ] ] + ; + tac2rec_fieldexprs: + [ [ f = tac2rec_fieldexpr; ";"; l = tac2rec_fieldexprs -> f :: l + | f = tac2rec_fieldexpr; ";" -> [f] + | f = tac2rec_fieldexpr-> [f] + | -> [] ] ] + ; + tac2rec_fieldexpr: + [ [ qid = Prim.qualid; ":="; e = tac2expr LEVEL "1" -> RelId qid, e ] ] + ; + tac2typ_prm: + [ [ -> [] + | id = typ_param -> [Loc.tag ~loc:!@loc id] + | "("; ids = LIST1 [ id = typ_param -> Loc.tag ~loc:!@loc id ] SEP "," ;")" -> ids + ] ] + ; + tac2typ_def: + [ [ prm = tac2typ_prm; id = Prim.qualid; (r, e) = tac2type_body -> (id, r, (prm, e)) ] ] + ; + tac2type_body: + [ [ -> false, CTydDef None + | ":="; e = tac2typ_knd -> false, e + | "::="; e = tac2typ_knd -> true, e + ] ] + ; + tac2def_typ: + [ [ "Type"; isrec = rec_flag; l = LIST1 tac2typ_def SEP "with" -> + StrTyp (isrec, l) + ] ] + ; + tac2def_ext: + [ [ "@"; IDENT "external"; id = locident; ":"; t = tac2type LEVEL "5"; ":="; + plugin = Prim.string; name = Prim.string -> + let ml = { mltac_plugin = plugin; mltac_tactic = name } in + StrPrm (id, t, ml) + ] ] + ; + syn_node: + [ [ "_" -> Loc.tag ~loc:!@loc None + | id = Prim.ident -> Loc.tag ~loc:!@loc (Some id) + ] ] + ; + sexpr: + [ [ s = Prim.string -> SexprStr (Loc.tag ~loc:!@loc s) + | n = Prim.integer -> SexprInt (Loc.tag ~loc:!@loc n) + | id = syn_node -> SexprRec (!@loc, id, []) + | id = syn_node; "("; tok = LIST1 sexpr SEP "," ; ")" -> + SexprRec (!@loc, id, tok) + ] ] + ; + syn_level: + [ [ -> None + | ":"; n = Prim.integer -> Some n + ] ] + ; + tac2def_syn: + [ [ "Notation"; toks = LIST1 sexpr; n = syn_level; ":="; + e = tac2expr -> + StrSyn (toks, n, e) + ] ] + ; +END + +GEXTEND Gram + Pcoq.Constr.operconstr: LEVEL "0" + [ [ IDENT "ltac2"; ":"; "("; tac = tac2expr; ")" -> + let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in + CAst.make ~loc:!@loc (CHole (None, IntroAnonymous, Some arg)) ] ] + ; +END + +let pr_ltac2entry _ = mt () (** FIXME *) +let pr_ltac2expr _ = mt () (** FIXME *) + +VERNAC ARGUMENT EXTEND ltac2_entry +PRINTED BY pr_ltac2entry +| [ tac2def_val(v) ] -> [ v ] +| [ tac2def_typ(t) ] -> [ t ] +| [ tac2def_ext(e) ] -> [ e ] +| [ tac2def_syn(e) ] -> [ e ] +END + +VERNAC COMMAND EXTEND VernacDeclareTactic2Definition CLASSIFIED AS SIDEFF +| [ "Ltac2" ltac2_entry(e) ] -> [ + let local = Locality.LocalityFixme.consume () in + Tac2entries.register_struct ?local e + ] +END + +let _ = + let mode = { + Proof_global.name = "Ltac2"; + set = (fun () -> set_command_entry tac2mode); + reset = (fun () -> set_command_entry Pcoq.Vernac_.noedit_mode); + } in + Proof_global.register_proof_mode mode + +VERNAC ARGUMENT EXTEND ltac2_expr +PRINTED BY pr_ltac2expr +| [ tac2expr(e) ] -> [ e ] +END + +open G_ltac +open Vernac_classifier + +VERNAC tac2mode EXTEND VernacLtac2 +| [ - ltac2_expr(t) ltac_use_default(default) ] => + [ classify_as_proofstep ] -> [ +(* let g = Option.default (Proof_global.get_default_goal_selector ()) g in *) + Tac2entries.call ~default t + ] +END + +open Stdarg + +VERNAC COMMAND EXTEND Ltac2Print CLASSIFIED AS SIDEFF +| [ "Print" "Ltac2" reference(tac) ] -> [ Tac2entries.print_ltac tac ] +END + diff --git a/src/ltac2_plugin.mlpack b/src/ltac2_plugin.mlpack new file mode 100644 index 0000000000..3d87a8cddb --- /dev/null +++ b/src/ltac2_plugin.mlpack @@ -0,0 +1,7 @@ +Tac2env +Tac2print +Tac2intern +Tac2interp +Tac2entries +Tac2core +G_ltac2 diff --git a/src/tac2core.ml b/src/tac2core.ml new file mode 100644 index 0000000000..91a3bfa168 --- /dev/null +++ b/src/tac2core.ml @@ -0,0 +1,648 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* t +| _ -> assert false + +let val_constr = val_tag (topwit Stdarg.wit_constr) +let val_ident = val_tag (topwit Stdarg.wit_ident) +let val_pp = Val.create "ltac2:pp" + +let extract_val (type a) (tag : a Val.typ) (Val.Dyn (tag', v)) : a = +match Val.eq tag tag' with +| None -> assert false +| Some Refl -> v + +module Core = +struct + +let t_int = coq_core "int" +let t_string = coq_core "string" +let t_array = coq_core "array" +let t_unit = coq_core "unit" +let t_list = coq_core "list" +let t_constr = coq_core "constr" +let t_ident = coq_core "ident" +let t_option = coq_core "option" + +let c_nil = coq_core "[]" +let c_cons = coq_core "::" + +let c_none = coq_core "None" +let c_some = coq_core "Some" + +end + +open Core + +let v_unit = ValInt 0 +let v_nil = ValInt 0 +let v_cons v vl = ValBlk (0, [|v; vl|]) + +module Value = +struct + +let of_unit () = v_unit + +let to_unit = function +| ValInt 0 -> () +| _ -> assert false + +let of_int n = ValInt n +let to_int = function +| ValInt n -> n +| _ -> assert false + +let of_bool b = if b then ValInt 0 else ValInt 1 + +let to_bool = function +| ValInt 0 -> true +| ValInt 1 -> false +| _ -> assert false + +let of_char n = ValInt (Char.code n) +let to_char = function +| ValInt n -> Char.chr n +| _ -> assert false + +let of_string s = ValStr s +let to_string = function +| ValStr s -> s +| _ -> assert false + +let rec of_list = function +| [] -> v_nil +| x :: l -> v_cons x (of_list l) + +let rec to_list = function +| ValInt 0 -> [] +| ValBlk (0, [|v; vl|]) -> v :: to_list vl +| _ -> assert false + +let of_ext tag c = + ValExt (Val.Dyn (tag, c)) + +let to_ext tag = function +| ValExt e -> extract_val tag e +| _ -> assert false + +let of_constr c = of_ext val_constr c +let to_constr c = to_ext val_constr c + +let of_ident c = of_ext val_ident c +let to_ident c = to_ext val_ident c + +(** FIXME: handle backtrace in Ltac2 exceptions *) +let of_exn c = match fst c with +| LtacError (kn, c) -> ValOpn (kn, c) +| _ -> of_ext val_exn c + +let to_exn c = match c with +| ValOpn (kn, c) -> (LtacError (kn, c), Exninfo.null) +| _ -> to_ext val_exn c + +let of_pp c = of_ext val_pp c +let to_pp c = to_ext val_pp c + +end + +let val_valexpr = Val.create "ltac2:valexpr" + +(** Stdlib exceptions *) + +let err_notfocussed = + LtacError (coq_core "Not_focussed", [||]) + +let err_outofbounds = + LtacError (coq_core "Out_of_bounds", [||]) + +let err_notfound = + LtacError (coq_core "Not_found", [||]) + +(** Helper functions *) + +let thaw f = interp_app f [v_unit] +let throw e = Proofview.tclLIFT (Proofview.NonLogical.raise e) + +let return x = Proofview.tclUNIT x +let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } + +let wrap f = + return () >>= fun () -> return (f ()) + +let wrap_unit f = + return () >>= fun () -> f (); return v_unit + +let pf_apply f = + Proofview.Goal.goals >>= function + | [] -> + Proofview.tclENV >>= fun env -> + Proofview.tclEVARMAP >>= fun sigma -> + f env sigma + | [gl] -> + gl >>= fun gl -> + f (Proofview.Goal.env gl) (Tacmach.New.project gl) + | _ :: _ :: _ -> + throw err_notfocussed + +(** Primitives *) + +(** Printing *) + +let prm_print : ml_tactic = function +| [pp] -> wrap_unit (fun () -> Feedback.msg_notice (Value.to_pp pp)) +| _ -> assert false + +let prm_message_of_int : ml_tactic = function +| [ValInt s] -> return (ValExt (Val.Dyn (val_pp, int s))) +| _ -> assert false + +let prm_message_of_string : ml_tactic = function +| [ValStr s] -> return (ValExt (Val.Dyn (val_pp, str (Bytes.to_string s)))) +| _ -> assert false + +let prm_message_of_constr : ml_tactic = function +| [c] -> + pf_apply begin fun env sigma -> + let c = Value.to_constr c in + let pp = Printer.pr_econstr_env env sigma c in + return (ValExt (Val.Dyn (val_pp, pp))) + end +| _ -> assert false + +let prm_message_concat : ml_tactic = function +| [m1; m2] -> + let m1 = Value.to_pp m1 in + let m2 = Value.to_pp m2 in + return (Value.of_pp (Pp.app m1 m2)) +| _ -> assert false + +(** Array *) + +let prm_array_make : ml_tactic = function +| [ValInt n; x] -> + if n < 0 || n > Sys.max_array_length then throw err_outofbounds + else wrap (fun () -> ValBlk (0, Array.make n x)) +| _ -> assert false + +let prm_array_length : ml_tactic = function +| [ValBlk (_, v)] -> return (ValInt (Array.length v)) +| _ -> assert false + +let prm_array_set : ml_tactic = function +| [ValBlk (_, v); ValInt n; x] -> + if n < 0 || n >= Array.length v then throw err_outofbounds + else wrap_unit (fun () -> v.(n) <- x) +| _ -> assert false + +let prm_array_get : ml_tactic = function +| [ValBlk (_, v); ValInt n] -> + if n < 0 || n >= Array.length v then throw err_outofbounds + else wrap (fun () -> v.(n)) +| _ -> assert false + +(** Int *) + +let prm_int_equal : ml_tactic = function +| [m; n] -> + return (Value.of_bool (Value.to_int m == Value.to_int n)) +| _ -> assert false + +let binop f : ml_tactic = function +| [m; n] -> return (Value.of_int (f (Value.to_int m) (Value.to_int n))) +| _ -> assert false + +let prm_int_compare args = binop Int.compare args +let prm_int_add args = binop (+) args +let prm_int_sub args = binop (-) args +let prm_int_mul args = binop ( * ) args + +let prm_int_neg : ml_tactic = function +| [m] -> return (Value.of_int (~- (Value.to_int m))) +| _ -> assert false + +(** String *) + +let prm_string_make : ml_tactic = function +| [n; c] -> + let n = Value.to_int n in + let c = Value.to_char c in + if n < 0 || n > Sys.max_string_length then throw err_outofbounds + else wrap (fun () -> Value.of_string (Bytes.make n c)) +| _ -> assert false + +let prm_string_length : ml_tactic = function +| [s] -> + return (Value.of_int (Bytes.length (Value.to_string s))) +| _ -> assert false + +let prm_string_set : ml_tactic = function +| [s; n; c] -> + let s = Value.to_string s in + let n = Value.to_int n in + let c = Value.to_char c in + if n < 0 || n >= Bytes.length s then throw err_outofbounds + else wrap_unit (fun () -> Bytes.set s n c) +| _ -> assert false + +let prm_string_get : ml_tactic = function +| [s; n] -> + let s = Value.to_string s in + let n = Value.to_int n in + if n < 0 || n >= Bytes.length s then throw err_outofbounds + else wrap (fun () -> Value.of_char (Bytes.get s n)) +| _ -> assert false + +(** Terms *) + +(** constr -> constr *) +let prm_constr_type : ml_tactic = function +| [c] -> + let c = Value.to_constr c in + let get_type env sigma = + Proofview.V82.wrap_exceptions begin fun () -> + let (sigma, t) = Typing.type_of env sigma c in + let t = Value.of_constr t in + Proofview.Unsafe.tclEVARS sigma <*> Proofview.tclUNIT t + end in + pf_apply get_type +| _ -> assert false + +(** constr -> constr *) +let prm_constr_equal : ml_tactic = function +| [c1; c2] -> + let c1 = Value.to_constr c1 in + let c2 = Value.to_constr c2 in + Proofview.tclEVARMAP >>= fun sigma -> + let b = EConstr.eq_constr sigma c1 c2 in + Proofview.tclUNIT (Value.of_bool b) +| _ -> assert false + +(** Error *) + +let prm_throw : ml_tactic = function +| [e] -> + let (e, info) = Value.to_exn e in + Proofview.tclLIFT (Proofview.NonLogical.raise ~info e) +| _ -> assert false + +(** Control *) + +(** exn -> 'a *) +let prm_zero : ml_tactic = function +| [e] -> + let (e, info) = Value.to_exn e in + Proofview.tclZERO ~info e +| _ -> assert false + +(** exn -> 'a *) +let prm_plus : ml_tactic = function +| [x; k] -> + Proofview.tclOR (thaw x) (fun e -> interp_app k [Value.of_exn e]) +| _ -> assert false + +(** (unit -> 'a) -> 'a *) +let prm_once : ml_tactic = function +| [f] -> Proofview.tclONCE (thaw f) +| _ -> assert false + +(** (unit -> unit) list -> unit *) +let prm_dispatch : ml_tactic = function +| [l] -> + let l = Value.to_list l in + let l = List.map (fun f -> Proofview.tclIGNORE (thaw f)) l in + Proofview.tclDISPATCH l >>= fun () -> return v_unit +| _ -> assert false + +(** (unit -> unit) list -> (unit -> unit) -> (unit -> unit) list -> unit *) +let prm_extend : ml_tactic = function +| [lft; tac; rgt] -> + let lft = Value.to_list lft in + let lft = List.map (fun f -> Proofview.tclIGNORE (thaw f)) lft in + let tac = Proofview.tclIGNORE (thaw tac) in + let rgt = Value.to_list rgt in + let rgt = List.map (fun f -> Proofview.tclIGNORE (thaw f)) rgt in + Proofview.tclEXTEND lft tac rgt >>= fun () -> return v_unit +| _ -> assert false + +(** (unit -> unit) -> unit *) +let prm_enter : ml_tactic = function +| [f] -> + let f = Proofview.tclIGNORE (thaw f) in + Proofview.tclINDEPENDENT f >>= fun () -> return v_unit +| _ -> assert false + +(** int -> int -> (unit -> 'a) -> 'a *) +let prm_focus : ml_tactic = function +| [i; j; tac] -> + let i = Value.to_int i in + let j = Value.to_int j in + Proofview.tclFOCUS i j (thaw tac) +| _ -> assert false + +(** unit -> unit *) +let prm_shelve : ml_tactic = function +| [_] -> Proofview.shelve >>= fun () -> return v_unit +| _ -> assert false + +(** unit -> unit *) +let prm_shelve_unifiable : ml_tactic = function +| [_] -> Proofview.shelve_unifiable >>= fun () -> return v_unit +| _ -> assert false + +let prm_new_goal : ml_tactic = function +| [ev] -> + let ev = Evar.unsafe_of_int (Value.to_int ev) in + Proofview.tclEVARMAP >>= fun sigma -> + if Evd.mem sigma ev then + Proofview.Unsafe.tclNEWGOALS [ev] <*> Proofview.tclUNIT v_unit + else throw err_notfound +| _ -> assert false + +(** unit -> constr *) +let prm_goal : ml_tactic = function +| [_] -> + Proofview.Goal.enter_one begin fun gl -> + let concl = Tacmach.New.pf_nf_concl gl in + return (Value.of_constr concl) + end +| _ -> assert false + +(** ident -> constr *) +let prm_hyp : ml_tactic = function +| [id] -> + let id = Value.to_ident id in + pf_apply begin fun env _ -> + let mem = try ignore (Environ.lookup_named id env); true with Not_found -> false in + if mem then return (Value.of_constr (EConstr.mkVar id)) + else Tacticals.New.tclZEROMSG + (str "Hypothesis " ++ quote (Id.print id) ++ str " not found") (** FIXME: Do something more sensible *) + end +| _ -> assert false + +(** (unit -> constr) -> unit *) +let prm_refine : ml_tactic = function +| [c] -> + let c = thaw c >>= fun c -> Proofview.tclUNIT ((), Value.to_constr c) in + Proofview.Goal.nf_enter begin fun gl -> + Refine.generic_refine ~typecheck:true c gl + end >>= fun () -> return v_unit +| _ -> assert false + + +(** Registering *) + +let () = Tac2env.define_primitive (pname "print") prm_print +let () = Tac2env.define_primitive (pname "message_of_string") prm_message_of_string +let () = Tac2env.define_primitive (pname "message_of_int") prm_message_of_int +let () = Tac2env.define_primitive (pname "message_of_constr") prm_message_of_constr +let () = Tac2env.define_primitive (pname "message_concat") prm_message_concat + +let () = Tac2env.define_primitive (pname "array_make") prm_array_make +let () = Tac2env.define_primitive (pname "array_length") prm_array_length +let () = Tac2env.define_primitive (pname "array_get") prm_array_get +let () = Tac2env.define_primitive (pname "array_set") prm_array_set + +let () = Tac2env.define_primitive (pname "string_make") prm_string_make +let () = Tac2env.define_primitive (pname "string_length") prm_string_length +let () = Tac2env.define_primitive (pname "string_get") prm_string_get +let () = Tac2env.define_primitive (pname "string_set") prm_string_set + +let () = Tac2env.define_primitive (pname "constr_type") prm_constr_type +let () = Tac2env.define_primitive (pname "constr_equal") prm_constr_equal + +let () = Tac2env.define_primitive (pname "int_equal") prm_int_equal +let () = Tac2env.define_primitive (pname "int_compare") prm_int_compare +let () = Tac2env.define_primitive (pname "int_neg") prm_int_neg +let () = Tac2env.define_primitive (pname "int_add") prm_int_add +let () = Tac2env.define_primitive (pname "int_sub") prm_int_sub +let () = Tac2env.define_primitive (pname "int_mul") prm_int_mul + +let () = Tac2env.define_primitive (pname "throw") prm_throw + +let () = Tac2env.define_primitive (pname "zero") prm_zero +let () = Tac2env.define_primitive (pname "plus") prm_plus +let () = Tac2env.define_primitive (pname "once") prm_once +let () = Tac2env.define_primitive (pname "dispatch") prm_dispatch +let () = Tac2env.define_primitive (pname "extend") prm_extend +let () = Tac2env.define_primitive (pname "enter") prm_enter + +let () = Tac2env.define_primitive (pname "focus") prm_focus +let () = Tac2env.define_primitive (pname "shelve") prm_shelve +let () = Tac2env.define_primitive (pname "shelve_unifiable") prm_shelve_unifiable +let () = Tac2env.define_primitive (pname "new_goal") prm_new_goal +let () = Tac2env.define_primitive (pname "goal") prm_goal +let () = Tac2env.define_primitive (pname "hyp") prm_hyp +let () = Tac2env.define_primitive (pname "refine") prm_refine + +(** ML types *) + +let constr_flags () = + let open Pretyping in + { + use_typeclasses = true; + solve_unification_constraints = true; + use_hook = Pfedit.solve_by_implicit_tactic (); + fail_evar = true; + expand_evars = true + } + +let open_constr_no_classes_flags () = + let open Pretyping in + { + use_typeclasses = false; + solve_unification_constraints = true; + use_hook = Pfedit.solve_by_implicit_tactic (); + fail_evar = false; + expand_evars = true + } + +(** Embed all Ltac2 data into Values *) +let to_lvar ist = + let open Glob_ops in + let map e = Val.Dyn (val_valexpr, e) in + let lfun = Id.Map.map map ist in + { empty_lvar with Glob_term.ltac_genargs = lfun } + +let interp_constr flags ist (c, _) = + let open Pretyping in + pf_apply begin fun env sigma -> + Proofview.V82.wrap_exceptions begin fun () -> + let ist = to_lvar ist in + let (sigma, c) = understand_ltac flags env sigma ist WithoutTypeConstraint c in + let c = Val.Dyn (val_constr, c) in + Proofview.Unsafe.tclEVARS sigma >>= fun () -> + Proofview.tclUNIT c + end + end + +let () = + let interp ist c = interp_constr (constr_flags ()) ist c in + let obj = { + ml_type = t_constr; + ml_interp = interp; + } in + define_ml_object Stdarg.wit_constr obj + +let () = + let interp ist c = interp_constr (open_constr_no_classes_flags ()) ist c in + let obj = { + ml_type = t_constr; + ml_interp = interp; + } in + define_ml_object Stdarg.wit_open_constr obj + +let () = + let interp _ id = return (Val.Dyn (val_ident, id)) in + let obj = { + ml_type = t_ident; + ml_interp = interp; + } in + define_ml_object Stdarg.wit_ident obj + +let () = + let interp ist env sigma concl tac = + let fold id (Val.Dyn (tag, v)) (accu : environment) : environment = + match Val.eq tag val_valexpr with + | None -> accu + | Some Refl -> Id.Map.add id v accu + in + let ist = Id.Map.fold fold ist Id.Map.empty in + let tac = Proofview.tclIGNORE (interp ist tac) in + let c, sigma = Pfedit.refine_by_tactic env sigma concl tac in + (EConstr.of_constr c, sigma) + in + Pretyping.register_constr_interp0 wit_ltac2 interp + +(** Built-in notation scopes *) + +let add_scope s f = + Tac2entries.register_scope (Id.of_string s) f + +let scope_fail () = CErrors.user_err (str "Invalid parsing token") + +let dummy_loc = Loc.make_loc (-1, -1) + +let rthunk e = + let loc = Tac2intern.loc_of_tacexpr e in + let var = [Loc.tag ~loc Anonymous, Some (CTypRef (loc, AbsKn Core.t_unit, []))] in + CTacFun (loc, var, e) + +let add_generic_scope s entry arg = + let parse = function + | [] -> + let scope = Extend.Aentry entry in + let act x = rthunk (CTacExt (dummy_loc, in_gen (rawwit arg) x)) in + Tac2entries.ScopeRule (scope, act) + | _ -> scope_fail () + in + add_scope s parse + +let () = add_scope "list0" begin function +| [tok] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let scope = Extend.Alist0 scope in + let act l = + let l = List.map act l in + CTacLst (None, l) + in + Tac2entries.ScopeRule (scope, act) +| [tok; SexprStr (_, str)] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let sep = Extend.Atoken (CLexer.terminal str) in + let scope = Extend.Alist0sep (scope, sep) in + let act l = + let l = List.map act l in + CTacLst (None, l) + in + Tac2entries.ScopeRule (scope, act) +| _ -> scope_fail () +end + +let () = add_scope "list1" begin function +| [tok] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let scope = Extend.Alist1 scope in + let act l = + let l = List.map act l in + CTacLst (None, l) + in + Tac2entries.ScopeRule (scope, act) +| [tok; SexprStr (_, str)] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let sep = Extend.Atoken (CLexer.terminal str) in + let scope = Extend.Alist1sep (scope, sep) in + let act l = + let l = List.map act l in + CTacLst (None, l) + in + Tac2entries.ScopeRule (scope, act) +| _ -> scope_fail () +end + +let () = add_scope "opt" begin function +| [tok] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let scope = Extend.Aopt scope in + let act opt = match opt with + | None -> + CTacRef (AbsKn (TacConstructor Core.c_none)) + | Some x -> + CTacApp (dummy_loc, CTacRef (AbsKn (TacConstructor Core.c_some)), [act x]) + in + Tac2entries.ScopeRule (scope, act) +| _ -> scope_fail () +end + +let () = add_scope "self" begin function +| [] -> + let scope = Extend.Aself in + let act tac = rthunk tac in + Tac2entries.ScopeRule (scope, act) +| _ -> scope_fail () +end + +let () = add_scope "next" begin function +| [] -> + let scope = Extend.Anext in + let act tac = rthunk tac in + Tac2entries.ScopeRule (scope, act) +| _ -> scope_fail () +end + +let () = add_scope "tactic" begin function +| [] -> + (** Default to level 5 parsing *) + let scope = Extend.Aentryl (Tac2entries.Pltac.tac2expr, 5) in + let act tac = rthunk tac in + Tac2entries.ScopeRule (scope, act) +| [SexprInt (loc, n)] -> + let () = if n < 0 || n > 5 then scope_fail () in + let scope = Extend.Aentryl (Tac2entries.Pltac.tac2expr, n) in + let act tac = rthunk tac in + Tac2entries.ScopeRule (scope, act) +| _ -> scope_fail () +end + +let () = add_generic_scope "ident" Pcoq.Prim.ident Stdarg.wit_ident +let () = add_generic_scope "constr" Pcoq.Constr.constr Stdarg.wit_constr diff --git a/src/tac2core.mli b/src/tac2core.mli new file mode 100644 index 0000000000..fc90499ac6 --- /dev/null +++ b/src/tac2core.mli @@ -0,0 +1,62 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* valexpr +val to_unit : valexpr -> unit + +val of_int : int -> valexpr +val to_int : valexpr -> int + +val of_bool : bool -> valexpr +val to_bool : valexpr -> bool + +val of_char : char -> valexpr +val to_char : valexpr -> char + +val of_list : valexpr list -> valexpr +val to_list : valexpr -> valexpr list + +val of_constr : EConstr.t -> valexpr +val to_constr : valexpr -> EConstr.t + +val of_exn : Exninfo.iexn -> valexpr +val to_exn : valexpr -> Exninfo.iexn + +val of_ident : Id.t -> valexpr +val to_ident : valexpr -> Id.t + +end diff --git a/src/tac2entries.ml b/src/tac2entries.ml new file mode 100644 index 0000000000..46f390a6d4 --- /dev/null +++ b/src/tac2entries.ml @@ -0,0 +1,648 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* obj = + declare_object {(default_object "TAC2-DEFINITION") with + cache_function = cache_tacdef; + load_function = load_tacdef; + open_function = open_tacdef; + subst_function = subst_tacdef; + classify_function = classify_tacdef} + +(** Type definition *) + +type typdef = { + typdef_local : bool; + typdef_expr : glb_quant_typedef; +} + +let change_kn_label kn id = + let (mp, dp, _) = KerName.repr kn in + KerName.make mp dp (Label.of_id id) + +let change_sp_label sp id = + let (dp, _) = Libnames.repr_path sp in + Libnames.make_path dp id + +let push_typedef visibility sp kn (_, def) = match def with +| GTydDef _ -> + Tac2env.push_type visibility sp kn +| GTydAlg cstrs -> + (** Register constructors *) + let iter (c, _) = + let spc = change_sp_label sp c in + let knc = change_kn_label kn c in + Tac2env.push_ltac visibility spc (TacConstructor knc) + in + Tac2env.push_type visibility sp kn; + List.iter iter cstrs +| GTydRec fields -> + (** Register fields *) + let iter (c, _, _) = + let spc = change_sp_label sp c in + let knc = change_kn_label kn c in + Tac2env.push_projection visibility spc knc + in + Tac2env.push_type visibility sp kn; + List.iter iter fields +| GTydOpn -> + Tac2env.push_type visibility sp kn + +let next i = + let ans = !i in + let () = incr i in + ans + +let define_typedef kn (params, def as qdef) = match def with +| GTydDef _ -> + Tac2env.define_type kn qdef +| GTydAlg cstrs -> + (** Define constructors *) + let constant = ref 0 in + let nonconstant = ref 0 in + let iter (c, args) = + let knc = change_kn_label kn c in + let tag = if List.is_empty args then next constant else next nonconstant in + let data = { + Tac2env.cdata_prms = params; + cdata_type = kn; + cdata_args = args; + cdata_indx = Some tag; + } in + Tac2env.define_constructor knc data + in + Tac2env.define_type kn qdef; + List.iter iter cstrs +| GTydRec fs -> + (** Define projections *) + let iter i (id, mut, t) = + let knp = change_kn_label kn id in + let proj = { + Tac2env.pdata_prms = params; + pdata_type = kn; + pdata_ptyp = t; + pdata_mutb = mut; + pdata_indx = i; + } in + Tac2env.define_projection knp proj + in + Tac2env.define_type kn qdef; + List.iteri iter fs +| GTydOpn -> + Tac2env.define_type kn qdef + +let perform_typdef vs ((sp, kn), def) = + let () = if not def.typdef_local then push_typedef vs sp kn def.typdef_expr in + define_typedef kn def.typdef_expr + +let load_typdef i obj = perform_typdef (Until i) obj +let open_typdef i obj = perform_typdef (Exactly i) obj + +let cache_typdef ((sp, kn), def) = + let () = push_typedef (Until 1) sp kn def.typdef_expr in + define_typedef kn def.typdef_expr + +let subst_typdef (subst, def) = + let expr' = subst_quant_typedef subst def.typdef_expr in + if expr' == def.typdef_expr then def else { def with typdef_expr = expr' } + +let classify_typdef o = Substitute o + +let inTypDef : typdef -> obj = + declare_object {(default_object "TAC2-TYPE-DEFINITION") with + cache_function = cache_typdef; + load_function = load_typdef; + open_function = open_typdef; + subst_function = subst_typdef; + classify_function = classify_typdef} + +(** Type extension *) + +type extension_data = { + edata_name : Id.t; + edata_args : int glb_typexpr list; +} + +type typext = { + typext_local : bool; + typext_prms : int; + typext_type : type_constant; + typext_expr : extension_data list; +} + +let push_typext vis sp kn def = + let iter data = + let spc = change_sp_label sp data.edata_name in + let knc = change_kn_label kn data.edata_name in + Tac2env.push_ltac vis spc (TacConstructor knc) + in + List.iter iter def.typext_expr + +let define_typext kn def = + let iter data = + let knc = change_kn_label kn data.edata_name in + let cdata = { + Tac2env.cdata_prms = def.typext_prms; + cdata_type = def.typext_type; + cdata_args = data.edata_args; + cdata_indx = None; + } in + Tac2env.define_constructor knc cdata + in + List.iter iter def.typext_expr + +let cache_typext ((sp, kn), def) = + let () = define_typext kn def in + push_typext (Until 1) sp kn def + +let perform_typext vs ((sp, kn), def) = + let () = if not def.typext_local then push_typext vs sp kn def in + define_typext kn def + +let load_typext i obj = perform_typext (Until i) obj +let open_typext i obj = perform_typext (Exactly i) obj + +let subst_typext (subst, e) = + let open Mod_subst in + let subst_data data = + let edata_args = List.smartmap (fun e -> subst_type subst e) data.edata_args in + if edata_args == data.edata_args then data + else { data with edata_args } + in + let typext_type = subst_kn subst e.typext_type in + let typext_expr = List.smartmap subst_data e.typext_expr in + if typext_type == e.typext_type && typext_expr == e.typext_expr then + e + else + { e with typext_type; typext_expr } + +let classify_typext o = Substitute o + +let inTypExt : typext -> obj = + declare_object {(default_object "TAC2-TYPE-EXTENSION") with + cache_function = cache_typext; + load_function = load_typext; + open_function = open_typext; + subst_function = subst_typext; + classify_function = classify_typext} + +(** Toplevel entries *) + +let dummy_loc = Loc.make_loc (-1, -1) + +let register_ltac ?(local = false) isrec tactics = + if isrec then + let map (na, e) = (na, None, e) in + let bindings = List.map map tactics in + let map ((loc, na), e) = match na with + | Anonymous -> None + | Name id -> + let qid = Libnames.qualid_of_ident id in + let e = CTacLet (dummy_loc, true, bindings, CTacRef (RelId (loc, qid))) in + let (e, t) = intern e in + let e = match e with + | GTacLet (true, _, e) -> assert false + | _ -> assert false + in + Some (e, t) + in + let tactics = List.map map tactics in + assert false (** FIXME *) + else + let map ((loc, na), e) = + let (e, t) = intern e in + let () = + if not (is_value e) then + user_err ?loc (str "Tactic definition must be a syntactical value") + in + let id = match na with + | Anonymous -> + user_err ?loc (str "Tactic definition must have a name") + | Name id -> id + in + let kn = Lib.make_kn id in + let exists = + try let _ = Tac2env.interp_global kn in true with Not_found -> false + in + let () = + if exists then + user_err ?loc (str "Tactic " ++ Nameops.pr_id id ++ str " already exists") + in + (id, e, t) + in + let defs = List.map map tactics in + let iter (id, e, t) = + let def = { + tacdef_local = local; + tacdef_expr = e; + tacdef_type = t; + } in + ignore (Lib.add_leaf id (inTacDef def)) + in + List.iter iter defs + +let qualid_to_ident (loc, qid) = + let (dp, id) = Libnames.repr_qualid qid in + if DirPath.is_empty dp then (loc, id) + else user_err ?loc (str "Identifier expected") + +let register_typedef ?(local = false) isrec types = + let same_name ((_, id1), _) ((_, id2), _) = Id.equal id1 id2 in + let () = match List.duplicates same_name types with + | [] -> () + | ((loc, id), _) :: _ -> + user_err ?loc (str "Multiple definition of the type name " ++ Id.print id) + in + let check ((loc, id), (params, def)) = + let same_name (_, id1) (_, id2) = Id.equal id1 id2 in + let () = match List.duplicates same_name params with + | [] -> () + | (loc, id) :: _ -> + user_err ?loc (str "The type parameter " ++ Id.print id ++ + str " occurs several times") + in + match def with + | CTydDef _ -> + if isrec then + user_err ?loc (str "The type abbreviation " ++ Id.print id ++ + str " cannot be recursive") + | CTydAlg cs -> + let same_name (id1, _) (id2, _) = Id.equal id1 id2 in + let () = match List.duplicates same_name cs with + | [] -> () + | (id, _) :: _ -> + user_err (str "Multiple definitions of the constructor " ++ Id.print id) + in + () + | CTydRec ps -> + let same_name (id1, _, _) (id2, _, _) = Id.equal id1 id2 in + let () = match List.duplicates same_name ps with + | [] -> () + | (id, _, _) :: _ -> + user_err (str "Multiple definitions of the projection " ++ Id.print id) + in + () + | CTydOpn -> + if isrec then + user_err ?loc (str "The open type declaration " ++ Id.print id ++ + str " cannot be recursive") + in + let () = List.iter check types in + let self = + if isrec then + let fold accu ((_, id), (params, _)) = + Id.Map.add id (Lib.make_kn id, List.length params) accu + in + List.fold_left fold Id.Map.empty types + else Id.Map.empty + in + let map ((_, id), def) = + let typdef = { + typdef_local = local; + typdef_expr = intern_typedef self def; + } in + (id, typdef) + in + let types = List.map map types in + let iter (id, def) = ignore (Lib.add_leaf id (inTypDef def)) in + List.iter iter types + +let register_primitive ?(local = false) (loc, id) t ml = + let t = intern_open_type t in + let rec count_arrow = function + | GTypArrow (_, t) -> 1 + count_arrow t + | _ -> 0 + in + let arrows = count_arrow (snd t) in + let () = if Int.equal arrows 0 then + user_err ?loc (str "External tactic must have at least one argument") in + let () = + try let _ = Tac2env.interp_primitive ml in () with Not_found -> + user_err ?loc (str "Unregistered primitive " ++ + quote (str ml.mltac_plugin) ++ spc () ++ quote (str ml.mltac_tactic)) + in + let init i = Id.of_string (Printf.sprintf "x%i" i) in + let names = List.init arrows init in + let bnd = List.map (fun id -> Name id) names in + let arg = List.map (fun id -> GTacVar id) names in + let e = GTacFun (bnd, GTacPrm (ml, arg)) in + let def = { + tacdef_local = local; + tacdef_expr = e; + tacdef_type = t; + } in + ignore (Lib.add_leaf id (inTacDef def)) + +let register_open ?(local = false) (loc, qid) (params, def) = + let kn = + try Tac2env.locate_type qid + with Not_found -> + user_err ?loc (str "Unbound type " ++ pr_qualid qid) + in + let (tparams, t) = Tac2env.interp_type kn in + let () = match t with + | GTydOpn -> () + | GTydAlg _ | GTydRec _ | GTydDef _ -> + user_err ?loc (str "Type " ++ pr_qualid qid ++ str " is not an open type") + in + let () = + let loc = Option.default dummy_loc loc in + if not (Int.equal (List.length params) tparams) then + Tac2intern.error_nparams_mismatch loc (List.length params) tparams + in + match def with + | CTydOpn -> () + | CTydAlg def -> + let intern_type t = + let tpe = CTydDef (Some t) in + let (_, ans) = intern_typedef Id.Map.empty (params, tpe) in + match ans with + | GTydDef (Some t) -> t + | _ -> assert false + in + let map (id, tpe) = + let tpe = List.map intern_type tpe in + { edata_name = id; edata_args = tpe } + in + let def = List.map map def in + let def = { + typext_local = local; + typext_type = kn; + typext_prms = tparams; + typext_expr = def; + } in + Lib.add_anonymous_leaf (inTypExt def) + | CTydRec _ | CTydDef _ -> + user_err ?loc (str "Extensions only accept inductive constructors") + +let register_type ?local isrec types = match types with +| [qid, true, def] -> + let (loc, _) = qid in + let () = if isrec then user_err ?loc (str "Extensions cannot be recursive") in + register_open ?local qid def +| _ -> + let map (qid, redef, def) = + let (loc, _) = qid in + let () = if redef then + user_err ?loc (str "Types can only be extended one by one") + in + (qualid_to_ident qid, def) + in + let types = List.map map types in + register_typedef ?local isrec types + +(** Parsing *) + +type 'a token = +| TacTerm of string +| TacNonTerm of Name.t * 'a + +type scope_rule = +| ScopeRule : (raw_tacexpr, 'a) Extend.symbol * ('a -> raw_tacexpr) -> scope_rule + +type scope_interpretation = sexpr list -> scope_rule + +let scope_table : scope_interpretation Id.Map.t ref = ref Id.Map.empty + +let register_scope id s = + scope_table := Id.Map.add id s !scope_table + +module ParseToken = +struct + +let loc_of_token = function +| SexprStr (loc, _) -> Option.default dummy_loc loc +| SexprInt (loc, _) -> Option.default dummy_loc loc +| SexprRec (loc, _, _) -> loc + +let parse_scope = function +| SexprRec (_, (loc, Some id), toks) -> + if Id.Map.mem id !scope_table then + Id.Map.find id !scope_table toks + else + CErrors.user_err ?loc (str "Unknown scope" ++ spc () ++ Nameops.pr_id id) +| tok -> + let loc = loc_of_token tok in + CErrors.user_err ~loc (str "Invalid parsing token") + +let parse_token = function +| SexprStr (_, s) -> TacTerm s +| SexprRec (_, (_, na), [tok]) -> + let na = match na with None -> Anonymous | Some id -> Name id in + let scope = parse_scope tok in + TacNonTerm (na, scope) +| tok -> + let loc = loc_of_token tok in + CErrors.user_err ~loc (str "Invalid parsing token") + +end + +let parse_scope = ParseToken.parse_scope + +type synext = { + synext_tok : sexpr list; + synext_exp : raw_tacexpr; + synext_lev : int option; + synext_loc : bool; +} + +type krule = +| KRule : + (raw_tacexpr, 'act, Loc.t -> raw_tacexpr) Extend.rule * + ((Loc.t -> (Name.t * raw_tacexpr) list -> raw_tacexpr) -> 'act) -> krule + +let rec get_rule (tok : scope_rule token list) : krule = match tok with +| [] -> KRule (Extend.Stop, fun k loc -> k loc []) +| TacNonTerm (na, ScopeRule (scope, inj)) :: tok -> + let KRule (rule, act) = get_rule tok in + let rule = Extend.Next (rule, scope) in + let act k e = act (fun loc acc -> k loc ((na, inj e) :: acc)) in + KRule (rule, act) +| TacTerm t :: tok -> + let KRule (rule, act) = get_rule tok in + let rule = Extend.Next (rule, Extend.Atoken (CLexer.terminal t)) in + let act k _ = act k in + KRule (rule, act) + +let perform_notation syn st = + let tok = List.rev_map ParseToken.parse_token syn.synext_tok in + let KRule (rule, act) = get_rule tok in + let mk loc args = + let map (na, e) = + let loc = loc_of_tacexpr e in + (Loc.tag ~loc na, None, e) + in + let bnd = List.map map args in + CTacLet (loc, false, bnd, syn.synext_exp) + in + let rule = Extend.Rule (rule, act mk) in + let lev = match syn.synext_lev with + | None -> None + | Some lev -> Some (string_of_int lev) + in + let rule = (lev, None, [rule]) in + ([Pcoq.ExtendRule (Pltac.tac2expr, None, (None, [rule]))], st) + +let ltac2_notation = + Pcoq.create_grammar_command "ltac2-notation" perform_notation + +let cache_synext (_, syn) = + Pcoq.extend_grammar_command ltac2_notation syn + +let open_synext i (_, syn) = + if Int.equal i 1 then Pcoq.extend_grammar_command ltac2_notation syn + +let subst_synext (subst, syn) = + let e = Tac2intern.subst_rawexpr subst syn.synext_exp in + if e == syn.synext_exp then syn else { syn with synext_exp = e } + +let classify_synext o = + if o.synext_loc then Dispose else Substitute o + +let inTac2Notation : synext -> obj = + declare_object {(default_object "TAC2-NOTATION") with + cache_function = cache_synext; + open_function = open_synext; + subst_function = subst_synext; + classify_function = classify_synext} + +let register_notation ?(local = false) tkn lev body = + (** Check that the tokens make sense *) + let entries = List.map ParseToken.parse_token tkn in + let fold accu tok = match tok with + | TacTerm _ -> accu + | TacNonTerm (Name id, _) -> Id.Set.add id accu + | TacNonTerm (Anonymous, _) -> accu + in + let ids = List.fold_left fold Id.Set.empty entries in + (** Globalize so that names are absolute *) + let body = Tac2intern.globalize ids body in + let ext = { + synext_tok = tkn; + synext_exp = body; + synext_lev = lev; + synext_loc = local; + } in + Lib.add_anonymous_leaf (inTac2Notation ext) + +(** Toplevel entries *) + +let register_struct ?local str = match str with +| StrVal (isrec, e) -> register_ltac ?local isrec e +| StrTyp (isrec, t) -> register_type ?local isrec t +| StrPrm (id, t, ml) -> register_primitive ?local id t ml +| StrSyn (tok, lev, e) -> register_notation ?local tok lev e + +(** Printing *) + +let print_ltac ref = + let (loc, qid) = qualid_of_reference ref in + let kn = + try Tac2env.locate_ltac qid + with Not_found -> user_err ?loc (str "Unknown tactic " ++ pr_qualid qid) + in + match kn with + | TacConstant kn -> + let (e, _, (_, t)) = Tac2env.interp_global kn in + let name = int_name () in + Feedback.msg_notice ( + hov 0 ( + hov 2 (pr_qualid qid ++ spc () ++ str ":" ++ spc () ++ pr_glbtype name t) ++ fnl () ++ + hov 2 (pr_qualid qid ++ spc () ++ str ":=" ++ spc () ++ pr_glbexpr e) + ) + ) + | TacConstructor kn -> + let _ = Tac2env.interp_constructor kn in + Feedback.msg_notice (hov 2 (str "Constructor" ++ spc () ++ str ":" ++ spc () ++ pr_qualid qid)) + +(** Calling tactics *) + +let solve default tac = + let status = Proof_global.with_current_proof begin fun etac p -> + let with_end_tac = if default then Some etac else None in + let (p, status) = Pfedit.solve SelectAll None tac ?with_end_tac p in + (* in case a strict subtree was completed, + go back to the top of the prooftree *) + let p = Proof.maximal_unfocus Vernacentries.command_focus p in + p, status + end in + if not status then Feedback.feedback Feedback.AddedAxiom + +let call ~default e = + let loc = loc_of_tacexpr e in + let (e, (_, t)) = intern e in + let () = check_unit ~loc t in + let tac = Tac2interp.interp Id.Map.empty e in + solve default (Proofview.tclIGNORE tac) + +(** Primitive algebraic types than can't be defined Coq-side *) + +let register_prim_alg name params def = + let id = Id.of_string name in + let def = List.map (fun (cstr, tpe) -> (Id.of_string_soft cstr, tpe)) def in + let def = (params, GTydAlg def) in + let def = { typdef_local = false; typdef_expr = def } in + ignore (Lib.add_leaf id (inTypDef def)) + +let coq_def n = KerName.make2 Tac2env.coq_prefix (Label.make n) + +let t_list = coq_def "list" + +let _ = Mltop.declare_cache_obj begin fun () -> + register_prim_alg "unit" 0 ["()", []]; + register_prim_alg "list" 1 [ + ("[]", []); + ("::", [GTypVar 0; GTypRef (t_list, [GTypVar 0])]); + ]; +end "ltac2_plugin" diff --git a/src/tac2entries.mli b/src/tac2entries.mli new file mode 100644 index 0000000000..71e8150057 --- /dev/null +++ b/src/tac2entries.mli @@ -0,0 +1,57 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* rec_flag -> + (Name.t located * raw_tacexpr) list -> unit + +val register_type : ?local:bool -> rec_flag -> + (qualid located * redef_flag * raw_quant_typedef) list -> unit + +val register_primitive : ?local:bool -> + Id.t located -> raw_typexpr -> ml_tactic_name -> unit + +val register_struct : ?local:bool -> strexpr -> unit + +val register_notation : ?local:bool -> sexpr list -> int option -> + raw_tacexpr -> unit + +(** {5 Notations} *) + +type scope_rule = +| ScopeRule : (raw_tacexpr, 'a) Extend.symbol * ('a -> raw_tacexpr) -> scope_rule + +type scope_interpretation = sexpr list -> scope_rule + +val register_scope : Id.t -> scope_interpretation -> unit +(** Create a new scope with the provided name *) + +val parse_scope : sexpr -> scope_rule +(** Use this to interpret the subscopes for interpretation functions *) + +(** {5 Inspecting} *) + +val print_ltac : Libnames.reference -> unit + +(** {5 Eval loop} *) + +(** Evaluate a tactic expression in the current environment *) +val call : default:bool -> raw_tacexpr -> unit + +(** {5 Parsing entries} *) + +module Pltac : +sig +val tac2expr : raw_tacexpr Pcoq.Gram.entry +end diff --git a/src/tac2env.ml b/src/tac2env.ml new file mode 100644 index 0000000000..5ccdd018ee --- /dev/null +++ b/src/tac2env.ml @@ -0,0 +1,242 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* ValInt n +| GTacRef kn -> + let (e, _) = + try KNmap.find kn ltac_state.contents.ltac_tactics + with Not_found -> assert false + in + eval_pure e +| GTacFun (na, e) -> + ValCls { clos_env = Id.Map.empty; clos_var = na; clos_exp = e } +| GTacCst (_, n, []) -> ValInt n +| GTacCst (_, n, el) -> ValBlk (n, Array.map_of_list eval_pure el) +| GTacOpn (kn, el) -> ValOpn (kn, Array.map_of_list eval_pure el) +| GTacAtm (AtmStr _) | GTacArr _ | GTacLet _ | GTacVar _ | GTacSet _ +| GTacApp _ | GTacCse _ | GTacPrj _ | GTacPrm _ | GTacExt _ | GTacWth _ -> + anomaly (Pp.str "Term is not a syntactical value") + +let define_global kn e = + let state = !ltac_state in + ltac_state := { state with ltac_tactics = KNmap.add kn e state.ltac_tactics } + +let interp_global kn = + let (e, t) = KNmap.find kn ltac_state.contents.ltac_tactics in + (e, eval_pure e, t) + +let define_constructor kn t = + let state = !ltac_state in + ltac_state := { state with ltac_constructors = KNmap.add kn t state.ltac_constructors } + +let interp_constructor kn = KNmap.find kn ltac_state.contents.ltac_constructors + +let define_projection kn t = + let state = !ltac_state in + ltac_state := { state with ltac_projections = KNmap.add kn t state.ltac_projections } + +let interp_projection kn = KNmap.find kn ltac_state.contents.ltac_projections + +let define_type kn e = + let state = !ltac_state in + ltac_state := { state with ltac_types = KNmap.add kn e state.ltac_types } + +let interp_type kn = KNmap.find kn ltac_state.contents.ltac_types + +module ML = +struct + type t = ml_tactic_name + let compare n1 n2 = + let c = String.compare n1.mltac_plugin n2.mltac_plugin in + if Int.equal c 0 then String.compare n1.mltac_tactic n2.mltac_tactic + else c +end + +module MLMap = Map.Make(ML) + +let primitive_map = ref MLMap.empty + +let define_primitive name f = primitive_map := MLMap.add name f !primitive_map +let interp_primitive name = MLMap.find name !primitive_map + +(** Name management *) + +module FullPath = +struct + type t = full_path + let equal = eq_full_path + let to_string = string_of_path + let repr sp = + let dir,id = repr_path sp in + id, (DirPath.repr dir) +end + +type tacref = Tac2expr.tacref = +| TacConstant of ltac_constant +| TacConstructor of ltac_constructor + +module TacRef = +struct +type t = tacref +let equal r1 r2 = match r1, r2 with +| TacConstant c1, TacConstant c2 -> KerName.equal c1 c2 +| TacConstructor c1, TacConstructor c2 -> KerName.equal c1 c2 +| _ -> false +end + +module KnTab = Nametab.Make(FullPath)(KerName) +module RfTab = Nametab.Make(FullPath)(TacRef) + +type nametab = { + tab_ltac : RfTab.t; + tab_ltac_rev : full_path KNmap.t * full_path KNmap.t; + tab_type : KnTab.t; + tab_type_rev : full_path KNmap.t; + tab_proj : KnTab.t; + tab_proj_rev : full_path KNmap.t; +} + +let empty_nametab = { + tab_ltac = RfTab.empty; + tab_ltac_rev = (KNmap.empty, KNmap.empty); + tab_type = KnTab.empty; + tab_type_rev = KNmap.empty; + tab_proj = KnTab.empty; + tab_proj_rev = KNmap.empty; +} + +let nametab = Summary.ref empty_nametab ~name:"ltac2-nametab" + +let push_ltac vis sp kn = + let tab = !nametab in + let tab_ltac = RfTab.push vis sp kn tab.tab_ltac in + let (constant_map, constructor_map) = tab.tab_ltac_rev in + let tab_ltac_rev = match kn with + | TacConstant c -> (KNmap.add c sp constant_map, constructor_map) + | TacConstructor c -> (constant_map, KNmap.add c sp constructor_map) + in + nametab := { tab with tab_ltac; tab_ltac_rev } + +let locate_ltac qid = + let tab = !nametab in + RfTab.locate qid tab.tab_ltac + +let locate_extended_all_ltac qid = + let tab = !nametab in + RfTab.find_prefixes qid tab.tab_ltac + +let shortest_qualid_of_ltac kn = + let tab = !nametab in + let sp = match kn with + | TacConstant c -> KNmap.find c (fst tab.tab_ltac_rev) + | TacConstructor c -> KNmap.find c (snd tab.tab_ltac_rev) + in + RfTab.shortest_qualid Id.Set.empty sp tab.tab_ltac + +let push_type vis sp kn = + let tab = !nametab in + let tab_type = KnTab.push vis sp kn tab.tab_type in + let tab_type_rev = KNmap.add kn sp tab.tab_type_rev in + nametab := { tab with tab_type; tab_type_rev } + +let locate_type qid = + let tab = !nametab in + KnTab.locate qid tab.tab_type + +let locate_extended_all_type qid = + let tab = !nametab in + KnTab.find_prefixes qid tab.tab_type + +let shortest_qualid_of_type kn = + let tab = !nametab in + let sp = KNmap.find kn tab.tab_type_rev in + KnTab.shortest_qualid Id.Set.empty sp tab.tab_type + +let push_projection vis sp kn = + let tab = !nametab in + let tab_proj = KnTab.push vis sp kn tab.tab_proj in + let tab_proj_rev = KNmap.add kn sp tab.tab_proj_rev in + nametab := { tab with tab_proj; tab_proj_rev } + +let locate_projection qid = + let tab = !nametab in + KnTab.locate qid tab.tab_proj + +let locate_extended_all_projection qid = + let tab = !nametab in + KnTab.find_prefixes qid tab.tab_proj + +let shortest_qualid_of_projection kn = + let tab = !nametab in + let sp = KNmap.find kn tab.tab_proj_rev in + KnTab.shortest_qualid Id.Set.empty sp tab.tab_proj + +type 'a ml_object = { + ml_type : type_constant; + ml_interp : environment -> 'a -> Geninterp.Val.t Proofview.tactic; +} + +module MLTypeObj = +struct + type ('a, 'b, 'c) obj = 'b ml_object + let name = "ltac2_ml_type" + let default _ = None +end + +module MLType = Genarg.Register(MLTypeObj) + +let define_ml_object t tpe = MLType.register0 t tpe +let interp_ml_object t = MLType.obj t + +(** Absolute paths *) + +let coq_prefix = + MPfile (DirPath.make (List.map Id.of_string ["Init"; "ltac2"; "Coq"])) + +(** Generic arguments *) + +let wit_ltac2 = Genarg.make0 "ltac2" diff --git a/src/tac2env.mli b/src/tac2env.mli new file mode 100644 index 0000000000..c4b8c1e0ca --- /dev/null +++ b/src/tac2env.mli @@ -0,0 +1,106 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* (glb_tacexpr * type_scheme) -> unit +val interp_global : ltac_constant -> (glb_tacexpr * valexpr * type_scheme) + +(** {5 Toplevel definition of types} *) + +val define_type : type_constant -> glb_quant_typedef -> unit +val interp_type : type_constant -> glb_quant_typedef + +(** {5 Toplevel definition of algebraic constructors} *) + +type constructor_data = { + cdata_prms : int; + (** Type parameters *) + cdata_type : type_constant; + (** Inductive definition to which the constructor pertains *) + cdata_args : int glb_typexpr list; + (** Types of the constructor arguments *) + cdata_indx : int option; + (** Index of the constructor in the ADT. Numbering is duplicated between + argumentless and argument-using constructors, e.g. in type ['a option] + [None] and [Some] have both index 0. This field is empty whenever the + constructor is a member of an open type. *) +} + +val define_constructor : ltac_constructor -> constructor_data -> unit +val interp_constructor : ltac_constructor -> constructor_data + +(** {5 Toplevel definition of projections} *) + +type projection_data = { + pdata_prms : int; + (** Type parameters *) + pdata_type : type_constant; + (** Record definition to which the projection pertains *) + pdata_ptyp : int glb_typexpr; + (** Type of the projection *) + pdata_mutb : bool; + (** Whether the field is mutable *) + pdata_indx : int; + (** Index of the projection *) +} + +val define_projection : ltac_projection -> projection_data -> unit +val interp_projection : ltac_projection -> projection_data + +(** {5 Name management} *) + +val push_ltac : visibility -> full_path -> tacref -> unit +val locate_ltac : qualid -> tacref +val locate_extended_all_ltac : qualid -> tacref list +val shortest_qualid_of_ltac : tacref -> qualid + +val push_type : visibility -> full_path -> type_constant -> unit +val locate_type : qualid -> type_constant +val locate_extended_all_type : qualid -> type_constant list +val shortest_qualid_of_type : type_constant -> qualid + +val push_projection : visibility -> full_path -> ltac_projection -> unit +val locate_projection : qualid -> ltac_projection +val locate_extended_all_projection : qualid -> ltac_projection list +val shortest_qualid_of_projection : ltac_projection -> qualid + +(** {5 Toplevel definitions of ML tactics} *) + +(** This state is not part of the summary, contrarily to the ones above. It is + intended to be used from ML plugins to register ML-side functions. *) + +val define_primitive : ml_tactic_name -> ml_tactic -> unit +val interp_primitive : ml_tactic_name -> ml_tactic + +(** {5 ML primitive types} *) + +type 'a ml_object = { + ml_type : type_constant; + ml_interp : environment -> 'a -> Geninterp.Val.t Proofview.tactic; +} + +val define_ml_object : ('a, 'b, 'c) genarg_type -> 'b ml_object -> unit +val interp_ml_object : ('a, 'b, 'c) genarg_type -> 'b ml_object + +(** {5 Absolute paths} *) + +val coq_prefix : ModPath.t +(** Path where primitive datatypes are defined in Ltac2 plugin. *) + +(** {5 Generic arguments} *) + +val wit_ltac2 : (raw_tacexpr, glb_tacexpr, Util.Empty.t) genarg_type diff --git a/src/tac2expr.mli b/src/tac2expr.mli new file mode 100644 index 0000000000..acdad9bab4 --- /dev/null +++ b/src/tac2expr.mli @@ -0,0 +1,195 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* valexpr Proofview.tactic + +type environment = valexpr Id.Map.t diff --git a/src/tac2intern.ml b/src/tac2intern.ml new file mode 100644 index 0000000000..b63e6a0cd8 --- /dev/null +++ b/src/tac2intern.ml @@ -0,0 +1,1454 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* elt -> bool +val create : unit -> 'a t +val fresh : 'a t -> elt +val find : elt -> 'a t -> (elt * 'a option) +val union : elt -> elt -> 'a t -> unit +val set : elt -> 'a -> 'a t -> unit +module Map : +sig + type key = elt + type +'a t + val empty : 'a t + val add : key -> 'a -> 'a t -> 'a t + val mem : key -> 'a t -> bool + val find : key -> 'a t -> 'a + val exists : (key -> 'a -> bool) -> 'a t -> bool +end +end += +struct +type elt = int +let equal = Int.equal +module Map = Int.Map + +type 'a node = +| Canon of int * 'a option +| Equiv of elt + +type 'a t = { + mutable uf_data : 'a node array; + mutable uf_size : int; +} + +let resize p = + if Int.equal (Array.length p.uf_data) p.uf_size then begin + let nsize = 2 * p.uf_size + 1 in + let v = Array.make nsize (Equiv 0) in + Array.blit p.uf_data 0 v 0 (Array.length p.uf_data); + p.uf_data <- v; + end + +let create () = { uf_data = [||]; uf_size = 0 } + +let fresh p = + resize p; + let n = p.uf_size in + p.uf_data.(n) <- (Canon (1, None)); + p.uf_size <- n + 1; + n + +let rec lookup n p = + let node = Array.get p.uf_data n in + match node with + | Canon (size, v) -> n, size, v + | Equiv y -> + let ((z, _, _) as res) = lookup y p in + if not (Int.equal z y) then Array.set p.uf_data n (Equiv z); + res + +let find n p = + let (x, _, v) = lookup n p in (x, v) + +let union x y p = + let ((x, size1, _) as xcan) = lookup x p in + let ((y, size2, _) as ycan) = lookup y p in + let xcan, ycan = if size1 < size2 then xcan, ycan else ycan, xcan in + let x, _, xnode = xcan in + let y, _, ynode = ycan in + assert (Option.is_empty xnode); + assert (Option.is_empty ynode); + p.uf_data.(x) <- Equiv y; + p.uf_data.(y) <- Canon (size1 + size2, None) + +let set x v p = + let (x, s, v') = lookup x p in + assert (Option.is_empty v'); + p.uf_data.(x) <- Canon (s, Some v) + +end + +type mix_var = +| GVar of UF.elt +| LVar of int + +type mix_type_scheme = int * mix_var glb_typexpr + +type environment = { + env_var : mix_type_scheme Id.Map.t; + (** Type schemes of bound variables *) + env_cst : UF.elt glb_typexpr UF.t; + (** Unification state *) + env_als : UF.elt Id.Map.t ref; + (** Map user-facing type variables to unification variables *) + env_opn : bool; + (** Accept unbound type variables *) + env_rec : (KerName.t * int) Id.Map.t; + (** Recursive type definitions *) +} + +let empty_env () = { + env_var = Id.Map.empty; + env_cst = UF.create (); + env_als = ref Id.Map.empty; + env_opn = true; + env_rec = Id.Map.empty; +} + +let env_name env = + (** Generate names according to a provided environment *) + let mk num = + let base = num mod 26 in + let rem = num / 26 in + let name = String.make 1 (Char.chr (97 + base)) in + let suff = if Int.equal rem 0 then "" else string_of_int rem in + let name = name ^ suff in + name + in + let fold id elt acc = UF.Map.add elt (Id.to_string id) acc in + let vars = Id.Map.fold fold env.env_als.contents UF.Map.empty in + let vars = ref vars in + let rec fresh n = + let name = mk n in + if UF.Map.exists (fun _ name' -> String.equal name name') !vars then fresh (succ n) + else name + in + fun n -> + if UF.Map.mem n !vars then UF.Map.find n !vars + else + let ans = fresh 0 in + let () = vars := UF.Map.add n ans !vars in + ans + +let ltac2_env : environment Genintern.Store.field = + Genintern.Store.field () + +let fresh_id env = UF.fresh env.env_cst + +let get_alias (loc, id) env = + try Id.Map.find id env.env_als.contents + with Not_found -> + if env.env_opn then + let n = fresh_id env in + let () = env.env_als := Id.Map.add id n env.env_als.contents in + n + else user_err ?loc (str "Unbound type parameter " ++ Id.print id) + +let push_name id t env = match id with +| Anonymous -> env +| Name id -> { env with env_var = Id.Map.add id t env.env_var } + +let dummy_loc = Loc.make_loc (-1, -1) + +let loc_of_tacexpr = function +| CTacAtm (loc, _) -> Option.default dummy_loc loc +| CTacRef (RelId (loc, _)) -> Option.default dummy_loc loc +| CTacRef (AbsKn _) -> dummy_loc +| CTacFun (loc, _, _) -> loc +| CTacApp (loc, _, _) -> loc +| CTacLet (loc, _, _, _) -> loc +| CTacTup (loc, _) -> Option.default dummy_loc loc +| CTacArr (loc, _) -> Option.default dummy_loc loc +| CTacLst (loc, _) -> Option.default dummy_loc loc +| CTacCnv (loc, _, _) -> loc +| CTacSeq (loc, _, _) -> loc +| CTacCse (loc, _, _) -> loc +| CTacRec (loc, _) -> loc +| CTacPrj (loc, _, _) -> loc +| CTacSet (loc, _, _, _) -> loc +| CTacExt (loc, _) -> loc + +let loc_of_patexpr = function +| CPatAny loc -> loc +| CPatRef (loc, _, _) -> loc +| CPatTup (loc, _) -> Option.default dummy_loc loc + +let error_nargs_mismatch loc nargs nfound = + user_err ~loc (str "Constructor expects " ++ int nargs ++ + str " arguments, but is applied to " ++ int nfound ++ + str " arguments") + +let error_nparams_mismatch loc nargs nfound = + user_err ~loc (str "Type expects " ++ int nargs ++ + str " arguments, but is applied to " ++ int nfound ++ + str " arguments") + +let rec subst_type subst (t : 'a glb_typexpr) = match t with +| GTypVar id -> subst id +| GTypArrow (t1, t2) -> GTypArrow (subst_type subst t1, subst_type subst t2) +| GTypTuple tl -> GTypTuple (List.map (fun t -> subst_type subst t) tl) +| GTypRef (qid, args) -> + GTypRef (qid, List.map (fun t -> subst_type subst t) args) + +let rec intern_type env (t : raw_typexpr) : UF.elt glb_typexpr = match t with +| CTypVar (loc, Name id) -> GTypVar (get_alias (Loc.tag ?loc id) env) +| CTypVar (_, Anonymous) -> GTypVar (fresh_id env) +| CTypRef (loc, rel, args) -> + let (kn, nparams) = match rel with + | RelId (loc, qid) -> + let (dp, id) = repr_qualid qid in + if DirPath.is_empty dp && Id.Map.mem id env.env_rec then + Id.Map.find id env.env_rec + else + let kn = + try Tac2env.locate_type qid + with Not_found -> + user_err ?loc (str "Unbound type constructor " ++ pr_qualid qid) + in + let (nparams, _) = Tac2env.interp_type kn in + (kn, nparams) + | AbsKn kn -> + let (nparams, _) = Tac2env.interp_type kn in + (kn, nparams) + in + let nargs = List.length args in + let () = + if not (Int.equal nparams nargs) then + let loc, qid = match rel with + | RelId lid -> lid + | AbsKn kn -> Some loc, shortest_qualid_of_type kn + in + user_err ?loc (strbrk "The type constructor " ++ pr_qualid qid ++ + strbrk " expects " ++ int nparams ++ strbrk " argument(s), but is here \ + applied to " ++ int nargs ++ strbrk "argument(s)") + in + GTypRef (kn, List.map (fun t -> intern_type env t) args) +| CTypArrow (loc, t1, t2) -> GTypArrow (intern_type env t1, intern_type env t2) +| CTypTuple (loc, tl) -> GTypTuple (List.map (fun t -> intern_type env t) tl) + +let fresh_type_scheme env (t : type_scheme) : UF.elt glb_typexpr = + let (n, t) = t in + let subst = Array.init n (fun _ -> fresh_id env) in + let substf i = GTypVar subst.(i) in + subst_type substf t + +let fresh_mix_type_scheme env (t : mix_type_scheme) : UF.elt glb_typexpr = + let (n, t) = t in + let subst = Array.init n (fun _ -> fresh_id env) in + let substf = function + | LVar i -> GTypVar subst.(i) + | GVar n -> GTypVar n + in + subst_type substf t + +let fresh_reftype env (kn : KerName.t) = + let (n, _) = Tac2env.interp_type kn in + let subst = Array.init n (fun _ -> fresh_id env) in + let t = GTypRef (kn, Array.map_to_list (fun i -> GTypVar i) subst) in + (subst, t) + +(** First-order unification algorithm *) + +let is_unfoldable kn = match snd (Tac2env.interp_type kn) with +| GTydDef (Some _) -> true +| GTydDef None | GTydAlg _ | GTydRec _ | GTydOpn -> false + +let unfold env kn args = + let (nparams, def) = Tac2env.interp_type kn in + let def = match def with + | GTydDef (Some t) -> t + | _ -> assert false + in + let args = Array.of_list args in + let subst n = args.(n) in + subst_type subst def + +(** View function, allows to ensure head normal forms *) +let rec kind env t = match t with +| GTypVar id -> + let (id, v) = UF.find id env.env_cst in + begin match v with + | None -> GTypVar id + | Some t -> kind env t + end +| GTypRef (kn, tl) -> + if is_unfoldable kn then kind env (unfold env kn tl) else t +| GTypArrow _ | GTypTuple _ -> t + +exception Occur + +let rec occur_check env id t = match kind env t with +| GTypVar id' -> if UF.equal id id' then raise Occur +| GTypArrow (t1, t2) -> + let () = occur_check env id t1 in + occur_check env id t2 +| GTypTuple tl -> + List.iter (fun t -> occur_check env id t) tl +| GTypRef (kn, tl) -> + List.iter (fun t -> occur_check env id t) tl + +exception CannotUnify of UF.elt glb_typexpr * UF.elt glb_typexpr + +let unify_var env id t = match kind env t with +| GTypVar id' -> + if not (UF.equal id id') then UF.union id id' env.env_cst +| GTypArrow _ | GTypRef _ | GTypTuple _ -> + try + let () = occur_check env id t in + UF.set id t env.env_cst + with Occur -> raise (CannotUnify (GTypVar id, t)) + +let rec unify env t1 t2 = match kind env t1, kind env t2 with +| GTypVar id, t | t, GTypVar id -> + unify_var env id t +| GTypArrow (t1, u1), GTypArrow (t2, u2) -> + let () = unify env t1 t2 in + unify env u1 u2 +| GTypTuple tl1, GTypTuple tl2 -> + if Int.equal (List.length tl1) (List.length tl2) then + List.iter2 (fun t1 t2 -> unify env t1 t2) tl1 tl2 + else raise (CannotUnify (t1, t2)) +| GTypRef (kn1, tl1), GTypRef (kn2, tl2) -> + if KerName.equal kn1 kn2 then + List.iter2 (fun t1 t2 -> unify env t1 t2) tl1 tl2 + else raise (CannotUnify (t1, t2)) +| _ -> raise (CannotUnify (t1, t2)) + +let unify ?loc env t1 t2 = + try unify env t1 t2 + with CannotUnify (u1, u2) -> + let name = env_name env in + user_err ?loc (str "This expression has type " ++ pr_glbtype name t1 ++ + str " but an expression what expected of type " ++ pr_glbtype name t2) + +(** Term typing *) + +let is_pure_constructor kn = + match snd (Tac2env.interp_type kn) with + | GTydAlg _ | GTydOpn -> true + | GTydRec fields -> + let is_pure (_, mut, _) = not mut in + List.for_all is_pure fields + | GTydDef _ -> assert false (** Type definitions have no constructors *) + +let rec is_value = function +| GTacAtm (AtmInt _) | GTacVar _ | GTacRef _ | GTacFun _ -> true +| GTacAtm (AtmStr _) | GTacApp _ | GTacLet _ -> false +| GTacCst (GCaseTuple _, _, el) -> List.for_all is_value el +| GTacCst (_, _, []) -> true +| GTacOpn (_, el) -> List.for_all is_value el +| GTacCst (GCaseAlg kn, _, el) -> is_pure_constructor kn && List.for_all is_value el +| GTacArr _ | GTacCse _ | GTacPrj _ | GTacSet _ | GTacExt _ | GTacPrm _ +| GTacWth _ -> false + +let is_rec_rhs = function +| GTacFun _ -> true +| GTacAtm _ | GTacVar _ | GTacRef _ | GTacApp _ | GTacLet _ | GTacPrj _ +| GTacSet _ | GTacArr _ | GTacExt _ | GTacPrm _ | GTacCst _ +| GTacCse _ | GTacOpn _ | GTacWth _ -> false + +let rec fv_type f t accu = match t with +| GTypVar id -> f id accu +| GTypArrow (t1, t2) -> fv_type f t1 (fv_type f t2 accu) +| GTypTuple tl -> List.fold_left (fun accu t -> fv_type f t accu) accu tl +| GTypRef (kn, tl) -> List.fold_left (fun accu t -> fv_type f t accu) accu tl + +let fv_env env = + let rec f id accu = match UF.find id env.env_cst with + | id, None -> UF.Map.add id () accu + | _, Some t -> fv_type f t accu + in + let fold_var id (_, t) accu = + let fmix id accu = match id with + | LVar _ -> accu + | GVar id -> f id accu + in + fv_type fmix t accu + in + let fv_var = Id.Map.fold fold_var env.env_var UF.Map.empty in + let fold_als _ id accu = f id accu in + Id.Map.fold fold_als !(env.env_als) fv_var + +let abstract_var env (t : UF.elt glb_typexpr) : mix_type_scheme = + let fv = fv_env env in + let count = ref 0 in + let vars = ref UF.Map.empty in + let rec subst id = + let (id, t) = UF.find id env.env_cst in + match t with + | None -> + if UF.Map.mem id fv then GTypVar (GVar id) + else + begin try UF.Map.find id !vars + with Not_found -> + let n = !count in + let var = GTypVar (LVar n) in + let () = incr count in + let () = vars := UF.Map.add id var !vars in + var + end + | Some t -> subst_type subst t + in + let t = subst_type subst t in + (!count, t) + +let monomorphic (t : UF.elt glb_typexpr) : mix_type_scheme = + let subst id = GTypVar (GVar id) in + (0, subst_type subst t) + +let warn_not_unit = + CWarnings.create ~name:"not-unit" ~category:"ltac" + (fun () -> strbrk "The following expression should have type unit.") + +let warn_redundant_clause = + CWarnings.create ~name:"redundant-clause" ~category:"ltac" + (fun () -> strbrk "The following clause is redundant.") + +let check_elt_unit loc env t = + let maybe_unit = match kind env t with + | GTypVar _ -> true + | GTypArrow _ | GTypTuple _ -> false + | GTypRef (kn, _) -> KerName.equal kn t_unit + in + if not maybe_unit then warn_not_unit ~loc () + +let check_elt_empty loc env t = match kind env t with +| GTypVar _ -> + user_err ~loc (str "Cannot infer an empty type for this expression") +| GTypArrow _ | GTypTuple _ -> + let name = env_name env in + user_err ~loc (str "Type " ++ pr_glbtype name t ++ str " is not an empty type") +| GTypRef (kn, _) -> + let def = Tac2env.interp_type kn in + match def with + | _, GTydAlg [] -> kn + | _ -> + let name = env_name env in + user_err ~loc (str "Type " ++ pr_glbtype name t ++ str " is not an empty type") + +let check_unit ?loc t = + let maybe_unit = match t with + | GTypVar _ -> true + | GTypArrow _ | GTypTuple _ -> false + | GTypRef (kn, _) -> KerName.equal kn t_unit + in + if not maybe_unit then warn_not_unit ?loc () + +let check_redundant_clause = function +| [] -> () +| (p, _) :: _ -> warn_redundant_clause ~loc:(loc_of_patexpr p) () + +let get_variable0 mem var = match var with +| RelId (loc, qid) -> + let (dp, id) = repr_qualid qid in + if DirPath.is_empty dp && mem id then ArgVar (loc, id) + else + let kn = + try Tac2env.locate_ltac qid + with Not_found -> + CErrors.user_err ?loc (str "Unbound value " ++ pr_qualid qid) + in + ArgArg kn +| AbsKn kn -> ArgArg kn + +let get_variable env var = + let mem id = Id.Map.mem id env.env_var in + get_variable0 mem var + +let get_constructor env var = match var with +| RelId (loc, qid) -> + let c = try Some (Tac2env.locate_ltac qid) with Not_found -> None in + begin match c with + | Some (TacConstructor knc) -> + let kn = Tac2env.interp_constructor knc in + ArgArg (kn, knc) + | Some (TacConstant _) -> + CErrors.user_err ?loc (str "The term " ++ pr_qualid qid ++ + str " is not the constructor of an inductive type.") + | None -> + let (dp, id) = repr_qualid qid in + if DirPath.is_empty dp then ArgVar (loc, id) + else CErrors.user_err ?loc (str "Unbound constructor " ++ pr_qualid qid) + end +| AbsKn knc -> + let kn = Tac2env.interp_constructor knc in + ArgArg (kn, knc) + +let get_projection var = match var with +| RelId (loc, qid) -> + let kn = try Tac2env.locate_projection qid with Not_found -> + user_err ?loc (pr_qualid qid ++ str " is not a projection") + in + Tac2env.interp_projection kn +| AbsKn kn -> + Tac2env.interp_projection kn + +let intern_atm env = function +| AtmInt n -> (GTacAtm (AtmInt n), GTypRef (t_int, [])) +| AtmStr s -> (GTacAtm (AtmStr s), GTypRef (t_string, [])) + +let invalid_pattern ?loc kn t = + let pt = match t with + | GCaseAlg kn' -> pr_typref kn + | GCaseTuple n -> str "tuple" + in + user_err ?loc (str "Invalid pattern, expected a pattern for type " ++ + pr_typref kn ++ str ", found a pattern of type " ++ pt) (** FIXME *) + +(** Pattern view *) + +type glb_patexpr = +| GPatVar of Name.t +| GPatRef of ltac_constructor * glb_patexpr list +| GPatTup of glb_patexpr list + +let rec intern_patexpr env = function +| CPatAny _ -> GPatVar Anonymous +| CPatRef (_, qid, []) -> + begin match get_constructor env qid with + | ArgVar (_, id) -> GPatVar (Name id) + | ArgArg (_, kn) -> GPatRef (kn, []) + end +| CPatRef (_, qid, pl) -> + begin match get_constructor env qid with + | ArgVar (loc, id) -> + user_err ?loc (str "Unbound constructor " ++ Nameops.pr_id id) + | ArgArg (_, kn) -> GPatRef (kn, List.map (fun p -> intern_patexpr env p) pl) + end +| CPatTup (_, pl) -> + GPatTup (List.map (fun p -> intern_patexpr env p) pl) + +type pattern_kind = +| PKind_empty +| PKind_variant of type_constant +| PKind_open of type_constant +| PKind_tuple of int +| PKind_any + +let get_pattern_kind env pl = match pl with +| [] -> PKind_empty +| p :: pl -> + let rec get_kind (p, _) pl = match intern_patexpr env p with + | GPatVar _ -> + begin match pl with + | [] -> PKind_any + | p :: pl -> get_kind p pl + end + | GPatRef (kn, pl) -> + let data = Tac2env.interp_constructor kn in + if Option.is_empty data.cdata_indx then PKind_open data.cdata_type + else PKind_variant data.cdata_type + | GPatTup tp -> PKind_tuple (List.length tp) + in + get_kind p pl + +(** Internalization *) + +let is_constructor env qid = match get_variable env qid with +| ArgArg (TacConstructor _) -> true +| _ -> false + +let rec intern_rec env = function +| CTacAtm (_, atm) -> intern_atm env atm +| CTacRef qid as e -> + begin match get_variable env qid with + | ArgVar (_, id) -> + let sch = Id.Map.find id env.env_var in + (GTacVar id, fresh_mix_type_scheme env sch) + | ArgArg (TacConstant kn) -> + let (_, _, sch) = Tac2env.interp_global kn in + (GTacRef kn, fresh_type_scheme env sch) + | ArgArg (TacConstructor kn) -> + let loc = loc_of_tacexpr e in + intern_constructor env loc kn [] + end +| CTacFun (loc, bnd, e) -> + let fold (env, bnd, tl) ((_, na), t) = + let t = match t with + | None -> GTypVar (fresh_id env) + | Some t -> intern_type env t + in + let env = push_name na (monomorphic t) env in + (env, na :: bnd, t :: tl) + in + let (env, bnd, tl) = List.fold_left fold (env, [], []) bnd in + let bnd = List.rev bnd in + let (e, t) = intern_rec env e in + let t = List.fold_left (fun accu t -> GTypArrow (t, accu)) t tl in + (GTacFun (bnd, e), t) +| CTacApp (loc, CTacRef qid, args) as e when is_constructor env qid -> + let kn = match get_variable env qid with + | ArgArg (TacConstructor kn) -> kn + | _ -> assert false + in + let loc = loc_of_tacexpr e in + intern_constructor env loc kn args +| CTacApp (loc, f, args) -> + let (f, ft) = intern_rec env f in + let fold arg (args, t) = + let (arg, argt) = intern_rec env arg in + (arg :: args, GTypArrow (argt, t)) + in + let ret = GTypVar (fresh_id env) in + let (args, t) = List.fold_right fold args ([], ret) in + let () = unify ~loc env ft t in + (GTacApp (f, args), ret) +| CTacLet (loc, false, el, e) -> + let fold accu ((loc, na), _, _) = match na with + | Anonymous -> accu + | Name id -> + if Id.Set.mem id accu then + user_err ?loc (str "Variable " ++ Id.print id ++ str " is bound several \ + times in this matching") + else Id.Set.add id accu + in + let _ = List.fold_left fold Id.Set.empty el in + let fold ((loc, na), tc, e) (el, p) = + let (e, t) = intern_rec env e in + let () = match tc with + | None -> () + | Some tc -> + let tc = intern_type env tc in + unify ?loc env t tc + in + let t = if is_value e then abstract_var env t else monomorphic t in + ((na, e) :: el), ((na, t) :: p) + in + let (el, p) = List.fold_right fold el ([], []) in + let nenv = List.fold_left (fun accu (na, t) -> push_name na t env) env p in + let (e, t) = intern_rec nenv e in + (GTacLet (false, el, e), t) +| CTacLet (loc, true, el, e) -> + intern_let_rec env loc el e +| CTacTup (loc, []) -> + (GTacCst (GCaseAlg t_unit, 0, []), GTypRef (t_unit, [])) +| CTacTup (loc, el) -> + let fold e (el, tl) = + let (e, t) = intern_rec env e in + (e :: el, t :: tl) + in + let (el, tl) = List.fold_right fold el ([], []) in + (GTacCst (GCaseTuple (List.length el), 0, el), GTypTuple tl) +| CTacArr (loc, []) -> + let id = fresh_id env in + (GTacArr [], GTypRef (t_int, [GTypVar id])) +| CTacArr (loc, e0 :: el) -> + let (e0, t0) = intern_rec env e0 in + let fold e el = intern_rec_with_constraint env e t0 :: el in + let el = e0 :: List.fold_right fold el [] in + (GTacArr el, GTypRef (t_array, [t0])) +| CTacLst (loc, []) -> + let id = fresh_id env in + (c_nil, GTypRef (t_list, [GTypVar id])) +| CTacLst (loc, e0 :: el) -> + let (e0, t0) = intern_rec env e0 in + let fold e el = c_cons (intern_rec_with_constraint env e t0) el in + let el = c_cons e0 (List.fold_right fold el c_nil) in + (el, GTypRef (t_list, [t0])) +| CTacCnv (loc, e, tc) -> + let (e, t) = intern_rec env e in + let tc = intern_type env tc in + let () = unify ~loc env t tc in + (e, tc) +| CTacSeq (loc, e1, e2) -> + let (e1, t1) = intern_rec env e1 in + let (e2, t2) = intern_rec env e2 in + let () = check_elt_unit loc env t1 in + (GTacLet (false, [Anonymous, e1], e2), t2) +| CTacCse (loc, e, pl) -> + intern_case env loc e pl +| CTacRec (loc, fs) -> + intern_record env loc fs +| CTacPrj (loc, e, proj) -> + let pinfo = get_projection proj in + let loc = loc_of_tacexpr e in + let (e, t) = intern_rec env e in + let subst = Array.init pinfo.pdata_prms (fun _ -> fresh_id env) in + let params = Array.map_to_list (fun i -> GTypVar i) subst in + let exp = GTypRef (pinfo.pdata_type, params) in + let () = unify ~loc env t exp in + let substf i = GTypVar subst.(i) in + let ret = subst_type substf pinfo.pdata_ptyp in + (GTacPrj (pinfo.pdata_type, e, pinfo.pdata_indx), ret) +| CTacSet (loc, e, proj, r) -> + let pinfo = get_projection proj in + let () = + if not pinfo.pdata_mutb then + let loc = match proj with + | RelId (loc, _) -> loc + | AbsKn _ -> None + in + user_err ?loc (str "Field is not mutable") + in + let subst = Array.init pinfo.pdata_prms (fun _ -> fresh_id env) in + let params = Array.map_to_list (fun i -> GTypVar i) subst in + let exp = GTypRef (pinfo.pdata_type, params) in + let e = intern_rec_with_constraint env e exp in + let substf i = GTypVar subst.(i) in + let ret = subst_type substf pinfo.pdata_ptyp in + let r = intern_rec_with_constraint env r ret in + (GTacSet (pinfo.pdata_type, e, pinfo.pdata_indx, r), GTypRef (t_unit, [])) +| CTacExt (loc, ext) -> + let open Genintern in + let GenArg (Rawwit tag, _) = ext in + let tpe = interp_ml_object tag in + (** External objects do not have access to the named context because this is + not stable by dynamic semantics. *) + let genv = Global.env_of_context Environ.empty_named_context_val in + let ist = empty_glob_sign genv in + let ist = { ist with extra = Store.set ist.extra ltac2_env env } in + let (_, ext) = Flags.with_option Ltac_plugin.Tacintern.strict_check (fun () -> generic_intern ist ext) () in + (GTacExt ext, GTypRef (tpe.ml_type, [])) + +and intern_rec_with_constraint env e exp = + let loc = loc_of_tacexpr e in + let (e, t) = intern_rec env e in + let () = unify ~loc env t exp in + e + +and intern_let_rec env loc el e = + let fold accu ((loc, na), _, _) = match na with + | Anonymous -> accu + | Name id -> + if Id.Set.mem id accu then + user_err ?loc (str "Variable " ++ Id.print id ++ str " is bound several \ + times in this matching") + else Id.Set.add id accu + in + let _ = List.fold_left fold Id.Set.empty el in + let map env ((loc, na), t, e) = + let id = fresh_id env in + let env = push_name na (monomorphic (GTypVar id)) env in + (env, (loc, na, t, e, id)) + in + let (env, el) = List.fold_map map env el in + let fold (loc, na, tc, e, id) (el, tl) = + let loc_e = loc_of_tacexpr e in + let (e, t) = intern_rec env e in + let () = + if not (is_rec_rhs e) then + user_err ~loc:loc_e (str "This kind of expression is not allowed as \ + right-hand side of a recursive binding") + in + let () = unify ?loc env t (GTypVar id) in + let () = match tc with + | None -> () + | Some tc -> + let tc = intern_type env tc in + unify ?loc env t tc + in + ((na, e) :: el, t :: tl) + in + let (el, tl) = List.fold_right fold el ([], []) in + let (e, t) = intern_rec env e in + (GTacLet (true, el, e), t) + +(** For now, patterns recognized by the pattern-matching compiling are limited + to depth-one where leaves are either variables or catch-all *) +and intern_case env loc e pl = + let (e', t) = intern_rec env e in + let todo ~loc () = user_err ~loc (str "Pattern not handled yet") in + match get_pattern_kind env pl with + | PKind_any -> + let (pat, b) = List.hd pl in + let na = match intern_patexpr env pat with + | GPatVar na -> na + | _ -> assert false + in + let () = check_redundant_clause (List.tl pl) in + let env = push_name na (monomorphic t) env in + let (b, tb) = intern_rec env b in + (GTacLet (false, [na, e'], b), tb) + | PKind_empty -> + let kn = check_elt_empty loc env t in + let r = fresh_id env in + (GTacCse (e', GCaseAlg kn, [||], [||]), GTypVar r) + | PKind_tuple len -> + begin match pl with + | [] -> assert false + | [CPatTup (_, []), b] -> + let () = unify ~loc:(loc_of_tacexpr e) env t (GTypRef (t_unit, [])) in + let (b, tb) = intern_rec env b in + (GTacCse (e', GCaseAlg t_unit, [|b|], [||]), tb) + | [CPatTup (_, pl), b] -> + let map = function + | CPatAny _ -> Anonymous + | CPatRef (loc, qid, []) -> + begin match get_constructor env qid with + | ArgVar (_, id) -> Name id + | ArgArg _ -> todo ~loc () + end + | p -> todo ~loc:(loc_of_patexpr p) () + in + let ids = Array.map_of_list map pl in + let tc = GTypTuple (List.map (fun _ -> GTypVar (fresh_id env)) pl) in + let () = unify ~loc:(loc_of_tacexpr e) env t tc in + let (b, tb) = intern_rec env b in + (GTacCse (e', GCaseTuple len, [||], [|ids, b|]), tb) + | (p, _) :: _ -> todo ~loc:(loc_of_patexpr p) () + end + | PKind_variant kn -> + let subst, tc = fresh_reftype env kn in + let () = unify ~loc:(loc_of_tacexpr e) env t tc in + let (params, def) = Tac2env.interp_type kn in + let cstrs = match def with + | GTydAlg c -> c + | _ -> assert false + in + let count (const, nonconst) (c, args) = match args with + | [] -> (succ const, nonconst) + | _ :: _ -> (const, succ nonconst) + in + let nconst, nnonconst = List.fold_left count (0, 0) cstrs in + let const = Array.make nconst None in + let nonconst = Array.make nnonconst None in + let ret = GTypVar (fresh_id env) in + let rec intern_branch = function + | [] -> () + | (pat, br) :: rem -> + let tbr = match pat with + | CPatAny _ -> + let () = check_redundant_clause rem in + let (br', brT) = intern_rec env br in + (** Fill all remaining branches *) + let fill (ncst, narg) (_, args) = + if List.is_empty args then + let () = + if Option.is_empty const.(ncst) then const.(ncst) <- Some br' + in + (succ ncst, narg) + else + let () = + if Option.is_empty const.(narg) then + let ids = Array.map_of_list (fun _ -> Anonymous) args in + nonconst.(narg) <- Some (ids, br') + in + (ncst, succ narg) + in + let _ = List.fold_left fill (0, 0) cstrs in + brT + | CPatRef (loc, qid, args) -> + let data = match get_constructor env qid with + | ArgVar _ -> todo ~loc () + | ArgArg (data, _) -> + let () = + let kn' = data.cdata_type in + if not (KerName.equal kn kn') then + invalid_pattern ~loc kn (GCaseAlg kn') + in + data + in + let get_id = function + | CPatAny _ -> Anonymous + | CPatRef (loc, qid, []) -> + begin match get_constructor env qid with + | ArgVar (_, id) -> Name id + | ArgArg _ -> todo ~loc () + end + | p -> todo ~loc:(loc_of_patexpr p) () + in + let ids = List.map get_id args in + let nids = List.length ids in + let nargs = List.length data.cdata_args in + let () = + if not (Int.equal nids nargs) then error_nargs_mismatch loc nargs nids + in + let fold env id tpe = + (** Instantiate all arguments *) + let subst n = GTypVar subst.(n) in + let tpe = subst_type subst tpe in + push_name id (monomorphic tpe) env + in + let nenv = List.fold_left2 fold env ids data.cdata_args in + let (br', brT) = intern_rec nenv br in + let () = + let index = match data.cdata_indx with + | Some i -> i + | None -> assert false + in + if List.is_empty args then + if Option.is_empty const.(index) then const.(index) <- Some br' + else warn_redundant_clause ~loc () + else + let ids = Array.of_list ids in + if Option.is_empty nonconst.(index) then nonconst.(index) <- Some (ids, br') + else warn_redundant_clause ~loc () + in + brT + | CPatTup (loc, tup) -> + invalid_pattern ?loc kn (GCaseTuple (List.length tup)) + in + let () = unify ~loc:(loc_of_tacexpr br) env ret tbr in + intern_branch rem + in + let () = intern_branch pl in + let map = function + | None -> user_err ~loc (str "TODO: Unhandled match case") (** FIXME *) + | Some x -> x + in + let const = Array.map map const in + let nonconst = Array.map map nonconst in + let ce = GTacCse (e', GCaseAlg kn, const, nonconst) in + (ce, ret) + | PKind_open kn -> + let subst, tc = fresh_reftype env kn in + let () = unify ~loc:(loc_of_tacexpr e) env t tc in + let ret = GTypVar (fresh_id env) in + let rec intern_branch map = function + | [] -> + user_err ~loc (str "Missing default case") + | (pat, br) :: rem -> + match intern_patexpr env pat with + | GPatVar na -> + let () = check_redundant_clause rem in + let nenv = push_name na (monomorphic tc) env in + let br' = intern_rec_with_constraint nenv br ret in + let def = (na, br') in + (map, def) + | GPatRef (knc, args) -> + let get = function + | GPatVar na -> na + | GPatRef _ | GPatTup _ -> + user_err ~loc (str "TODO: Unhandled match case") (** FIXME *) + in + let loc = loc_of_patexpr pat in + let ids = List.map get args in + let data = Tac2env.interp_constructor knc in + let () = + if not (KerName.equal kn data.cdata_type) then + invalid_pattern ~loc kn (GCaseAlg data.cdata_type) + in + let nids = List.length ids in + let nargs = List.length data.cdata_args in + let () = + if not (Int.equal nids nargs) then error_nargs_mismatch loc nargs nids + in + let fold env id tpe = + (** Instantiate all arguments *) + let subst n = GTypVar subst.(n) in + let tpe = subst_type subst tpe in + push_name id (monomorphic tpe) env + in + let nenv = List.fold_left2 fold env ids data.cdata_args in + let br' = intern_rec_with_constraint nenv br ret in + let map = + if KNmap.mem knc map then + let () = warn_redundant_clause ~loc () in + map + else + KNmap.add knc (Anonymous, Array.of_list ids, br') map + in + intern_branch map rem + | GPatTup tup -> + invalid_pattern ~loc kn (GCaseTuple (List.length tup)) + in + let (map, def) = intern_branch KNmap.empty pl in + (GTacWth { opn_match = e'; opn_branch = map; opn_default = def }, ret) + +and intern_constructor env loc kn args = + let cstr = interp_constructor kn in + let nargs = List.length cstr.cdata_args in + if Int.equal nargs (List.length args) then + let subst = Array.init cstr.cdata_prms (fun _ -> fresh_id env) in + let substf i = GTypVar subst.(i) in + let types = List.map (fun t -> subst_type substf t) cstr.cdata_args in + let ans = GTypRef (cstr.cdata_type, List.init cstr.cdata_prms (fun i -> GTypVar subst.(i))) in + let map arg tpe = intern_rec_with_constraint env arg tpe in + let args = List.map2 map args types in + match cstr.cdata_indx with + | Some idx -> + (GTacCst (GCaseAlg cstr.cdata_type, idx, args), ans) + | None -> + (GTacOpn (kn, args), ans) + else + error_nargs_mismatch loc nargs (List.length args) + +and intern_record env loc fs = + let map (proj, e) = + let loc = match proj with + | RelId (loc, _) -> loc + | AbsKn _ -> None + in + let proj = get_projection proj in + (loc, proj, e) + in + let fs = List.map map fs in + let kn = match fs with + | [] -> user_err ~loc (str "Cannot infer the corresponding record type") + | (_, proj, _) :: _ -> proj.pdata_type + in + let params, typdef = match Tac2env.interp_type kn with + | n, GTydRec def -> n, def + | _ -> assert false + in + let subst = Array.init params (fun _ -> fresh_id env) in + (** Set the answer [args] imperatively *) + let args = Array.make (List.length typdef) None in + let iter (loc, pinfo, e) = + if KerName.equal kn pinfo.pdata_type then + let index = pinfo.pdata_indx in + match args.(index) with + | None -> + let exp = subst_type (fun i -> GTypVar subst.(i)) pinfo.pdata_ptyp in + let e = intern_rec_with_constraint env e exp in + args.(index) <- Some e + | Some _ -> + let (name, _, _) = List.nth typdef pinfo.pdata_indx in + user_err ?loc (str "Field " ++ Id.print name ++ str " is defined \ + several times") + else + user_err ?loc (str "Field " ++ (*KerName.print knp ++*) str " does not \ + pertain to record definition " ++ pr_typref pinfo.pdata_type) + in + let () = List.iter iter fs in + let () = match Array.findi (fun _ o -> Option.is_empty o) args with + | None -> () + | Some i -> + let (field, _, _) = List.nth typdef i in + user_err ~loc (str "Field " ++ Id.print field ++ str " is undefined") + in + let args = Array.map_to_list Option.get args in + let tparam = List.init params (fun i -> GTypVar subst.(i)) in + (GTacCst (GCaseAlg kn, 0, args), GTypRef (kn, tparam)) + +let normalize env (count, vars) (t : UF.elt glb_typexpr) = + let get_var id = + try UF.Map.find id !vars + with Not_found -> + let () = assert env.env_opn in + let n = GTypVar !count in + let () = incr count in + let () = vars := UF.Map.add id n !vars in + n + in + let rec subst id = match UF.find id env.env_cst with + | id, None -> get_var id + | _, Some t -> subst_type subst t + in + subst_type subst t + +let intern e = + let env = empty_env () in + let (e, t) = intern_rec env e in + let count = ref 0 in + let vars = ref UF.Map.empty in + let t = normalize env (count, vars) t in + (e, (!count, t)) + +let intern_typedef self (ids, t) : glb_quant_typedef = + let env = { (empty_env ()) with env_rec = self } in + (** Initialize type parameters *) + let map id = get_alias id env in + let ids = List.map map ids in + let count = ref (List.length ids) in + let vars = ref UF.Map.empty in + let iter n id = vars := UF.Map.add id (GTypVar n) !vars in + let () = List.iteri iter ids in + (** Do not accept unbound type variables *) + let env = { env with env_opn = false } in + let intern t = + let t = intern_type env t in + normalize env (count, vars) t + in + let count = !count in + match t with + | CTydDef None -> (count, GTydDef None) + | CTydDef (Some t) -> (count, GTydDef (Some (intern t))) + | CTydAlg constrs -> + let map (c, t) = (c, List.map intern t) in + let constrs = List.map map constrs in + (count, GTydAlg constrs) + | CTydRec fields -> + let map (c, mut, t) = (c, mut, intern t) in + let fields = List.map map fields in + (count, GTydRec fields) + | CTydOpn -> (count, GTydOpn) + +let intern_open_type t = + let env = empty_env () in + let t = intern_type env t in + let count = ref 0 in + let vars = ref UF.Map.empty in + let t = normalize env (count, vars) t in + (!count, t) + +(** Globalization *) + +let add_name accu = function +| Name id -> Id.Set.add id accu +| Anonymous -> accu + +let get_projection0 var = match var with +| RelId (loc, qid) -> + let kn = try Tac2env.locate_projection qid with Not_found -> + user_err ?loc (pr_qualid qid ++ str " is not a projection") + in + kn +| AbsKn kn -> kn + +let rec globalize ids e = match e with +| CTacAtm _ -> e +| CTacRef ref -> + let mem id = Id.Set.mem id ids in + begin match get_variable0 mem ref with + | ArgVar _ -> e + | ArgArg kn -> CTacRef (AbsKn kn) + end +| CTacFun (loc, bnd, e) -> + let fold accu ((_, na), _) = add_name accu na in + let ids = List.fold_left fold ids bnd in + let e = globalize ids e in + CTacFun (loc, bnd, e) +| CTacApp (loc, e, el) -> + let e = globalize ids e in + let el = List.map (fun e -> globalize ids e) el in + CTacApp (loc, e, el) +| CTacLet (loc, isrec, bnd, e) -> + let fold accu ((_, na), _, _) = add_name accu na in + let ext = List.fold_left fold Id.Set.empty bnd in + let eids = Id.Set.union ext ids in + let e = globalize eids e in + let map (qid, t, e) = + let ids = if isrec then eids else ids in + (qid, t, globalize ids e) + in + let bnd = List.map map bnd in + CTacLet (loc, isrec, bnd, e) +| CTacTup (loc, el) -> + let el = List.map (fun e -> globalize ids e) el in + CTacTup (loc, el) +| CTacArr (loc, el) -> + let el = List.map (fun e -> globalize ids e) el in + CTacArr (loc, el) +| CTacLst (loc, el) -> + let el = List.map (fun e -> globalize ids e) el in + CTacLst (loc, el) +| CTacCnv (loc, e, t) -> + let e = globalize ids e in + CTacCnv (loc, e, t) +| CTacSeq (loc, e1, e2) -> + let e1 = globalize ids e1 in + let e2 = globalize ids e2 in + CTacSeq (loc, e1, e2) +| CTacCse (loc, e, bl) -> + let e = globalize ids e in + let bl = List.map (fun b -> globalize_case ids b) bl in + CTacCse (loc, e, bl) +| CTacRec (loc, r) -> + let map (p, e) = + let p = get_projection0 p in + let e = globalize ids e in + (AbsKn p, e) + in + CTacRec (loc, List.map map r) +| CTacPrj (loc, e, p) -> + let e = globalize ids e in + let p = get_projection0 p in + CTacPrj (loc, e, AbsKn p) +| CTacSet (loc, e, p, e') -> + let e = globalize ids e in + let p = get_projection0 p in + let e' = globalize ids e' in + CTacSet (loc, e, AbsKn p, e') +| CTacExt (loc, arg) -> + let arg = pr_argument_type (genarg_tag arg) in + CErrors.user_err ~loc (str "Cannot globalize generic arguments of type" ++ spc () ++ arg) + +and globalize_case ids (p, e) = + (globalize_pattern ids p, globalize ids e) + +and globalize_pattern ids p = match p with +| CPatAny _ -> p +| CPatRef (loc, cst, pl) -> + let cst = match get_constructor () cst with + | ArgVar _ -> cst + | ArgArg (_, knc) -> AbsKn knc + in + let pl = List.map (fun p -> globalize_pattern ids p) pl in + CPatRef (loc, cst, pl) +| CPatTup (loc, pl) -> + let pl = List.map (fun p -> globalize_pattern ids p) pl in + CPatTup (loc, pl) + +(** Kernel substitution *) + +open Mod_subst + +let rec subst_type subst t = match t with +| GTypVar _ -> t +| GTypArrow (t1, t2) -> + let t1' = subst_type subst t1 in + let t2' = subst_type subst t2 in + if t1' == t1 && t2' == t2 then t + else GTypArrow (t1', t2') +| GTypTuple tl -> + let tl'= List.smartmap (fun t -> subst_type subst t) tl in + if tl' == tl then t else GTypTuple tl' +| GTypRef (kn, tl) -> + let kn' = subst_kn subst kn in + let tl' = List.smartmap (fun t -> subst_type subst t) tl in + if kn' == kn && tl' == tl then t else GTypRef (kn', tl') + +let subst_case_info subst ci = match ci with +| GCaseAlg kn -> + let kn' = subst_kn subst kn in + if kn' == kn then ci else GCaseAlg kn' +| GCaseTuple _ -> ci + +let rec subst_expr subst e = match e with +| GTacAtm _ | GTacVar _ | GTacPrm _ -> e +| GTacRef kn -> GTacRef (subst_kn subst kn) +| GTacFun (ids, e) -> GTacFun (ids, subst_expr subst e) +| GTacApp (f, args) -> + GTacApp (subst_expr subst f, List.map (fun e -> subst_expr subst e) args) +| GTacLet (r, bs, e) -> + let bs = List.map (fun (na, e) -> (na, subst_expr subst e)) bs in + GTacLet (r, bs, subst_expr subst e) +| GTacArr el -> + GTacArr (List.map (fun e -> subst_expr subst e) el) +| GTacCst (t, n, el) as e0 -> + let t' = match t with + | GCaseAlg kn -> + let kn' = subst_kn subst kn in + if kn' == kn then t else GCaseAlg kn' + | GCaseTuple _ -> t + in + let el' = List.smartmap (fun e -> subst_expr subst e) el in + if t' == t && el' == el then e0 else GTacCst (t', n, el') +| GTacCse (e, ci, cse0, cse1) -> + let cse0' = Array.map (fun e -> subst_expr subst e) cse0 in + let cse1' = Array.map (fun (ids, e) -> (ids, subst_expr subst e)) cse1 in + let ci' = subst_case_info subst ci in + GTacCse (subst_expr subst e, ci', cse0', cse1') +| GTacWth { opn_match = e; opn_branch = br; opn_default = (na, def) } as e0 -> + let e' = subst_expr subst e in + let def' = subst_expr subst def in + let fold kn (self, vars, p) accu = + let kn' = subst_kn subst kn in + let p' = subst_expr subst p in + if kn' == kn && p' == p then accu + else KNmap.add kn' (self, vars, p') (KNmap.remove kn accu) + in + let br' = KNmap.fold fold br br in + if e' == e && br' == br && def' == def then e0 + else GTacWth { opn_match = e'; opn_default = (na, def'); opn_branch = br' } +| GTacPrj (kn, e, p) as e0 -> + let kn' = subst_kn subst kn in + let e' = subst_expr subst e in + if kn' == kn && e' == e then e0 else GTacPrj (kn', e', p) +| GTacSet (kn, e, p, r) as e0 -> + let kn' = subst_kn subst kn in + let e' = subst_expr subst e in + let r' = subst_expr subst r in + if kn' == kn && e' == e && r' == r then e0 else GTacSet (kn', e', p, r') +| GTacExt ext -> + let ext' = Genintern.generic_substitute subst ext in + if ext' == ext then e else GTacExt ext' +| GTacOpn (kn, el) as e0 -> + let kn' = subst_kn subst kn in + let el' = List.smartmap (fun e -> subst_expr subst e) el in + if kn' == kn && el' == el then e0 else GTacOpn (kn', el') + +let subst_typedef subst e = match e with +| GTydDef t -> + let t' = Option.smartmap (fun t -> subst_type subst t) t in + if t' == t then e else GTydDef t' +| GTydAlg constrs -> + let map (c, tl as p) = + let tl' = List.smartmap (fun t -> subst_type subst t) tl in + if tl' == tl then p else (c, tl') + in + let constrs' = List.smartmap map constrs in + if constrs' == constrs then e else GTydAlg constrs' +| GTydRec fields -> + let map (c, mut, t as p) = + let t' = subst_type subst t in + if t' == t then p else (c, mut, t') + in + let fields' = List.smartmap map fields in + if fields' == fields then e else GTydRec fields' +| GTydOpn -> GTydOpn + +let subst_quant_typedef subst (prm, def as qdef) = + let def' = subst_typedef subst def in + if def' == def then qdef else (prm, def') + +let subst_type_scheme subst (prm, t as sch) = + let t' = subst_type subst t in + if t' == t then sch else (prm, t') + +let subst_or_relid subst ref = match ref with +| RelId _ -> ref +| AbsKn kn -> + let kn' = subst_kn subst kn in + if kn' == kn then ref else AbsKn kn' + +let rec subst_rawtype subst t = match t with +| CTypVar _ -> t +| CTypArrow (loc, t1, t2) -> + let t1' = subst_rawtype subst t1 in + let t2' = subst_rawtype subst t2 in + if t1' == t1 && t2' == t2 then t else CTypArrow (loc, t1', t2') +| CTypTuple (loc, tl) -> + let tl' = List.smartmap (fun t -> subst_rawtype subst t) tl in + if tl' == tl then t else CTypTuple (loc, tl') +| CTypRef (loc, ref, tl) -> + let ref' = subst_or_relid subst ref in + let tl' = List.smartmap (fun t -> subst_rawtype subst t) tl in + if ref' == ref && tl' == tl then t else CTypRef (loc, ref', tl') + +let subst_tacref subst ref = match ref with +| RelId _ -> ref +| AbsKn (TacConstant kn) -> + let kn' = subst_kn subst kn in + if kn' == kn then ref else AbsKn (TacConstant kn') +| AbsKn (TacConstructor kn) -> + let kn' = subst_kn subst kn in + if kn' == kn then ref else AbsKn (TacConstructor kn') + +let subst_projection subst prj = match prj with +| RelId _ -> prj +| AbsKn kn -> + let kn' = subst_kn subst kn in + if kn' == kn then prj else AbsKn kn' + +let rec subst_rawpattern subst p = match p with +| CPatAny _ -> p +| CPatRef (loc, c, pl) -> + let pl' = List.smartmap (fun p -> subst_rawpattern subst p) pl in + let c' = match c with + | RelId _ -> c + | AbsKn kn -> + let kn' = subst_kn subst kn in + if kn' == kn then c else AbsKn kn' + in + if pl' == pl && c' == c then p else CPatRef (loc, c', pl') +| CPatTup (loc, pl) -> + let pl' = List.smartmap (fun p -> subst_rawpattern subst p) pl in + if pl' == pl then p else CPatTup (loc, pl') + +(** Used for notations *) +let rec subst_rawexpr subst t = match t with +| CTacAtm _ -> t +| CTacRef ref -> + let ref' = subst_tacref subst ref in + if ref' == ref then t else CTacRef ref' +| CTacFun (loc, bnd, e) -> + let map (na, t as p) = + let t' = Option.smartmap (fun t -> subst_rawtype subst t) t in + if t' == t then p else (na, t') + in + let bnd' = List.smartmap map bnd in + let e' = subst_rawexpr subst e in + if bnd' == bnd && e' == e then t else CTacFun (loc, bnd', e') +| CTacApp (loc, e, el) -> + let e' = subst_rawexpr subst e in + let el' = List.smartmap (fun e -> subst_rawexpr subst e) el in + if e' == e && el' == el then t else CTacApp (loc, e', el') +| CTacLet (loc, isrec, bnd, e) -> + let map (na, t, e as p) = + let t' = Option.smartmap (fun t -> subst_rawtype subst t) t in + let e' = subst_rawexpr subst e in + if t' == t && e' == e then p else (na, t', e') + in + let bnd' = List.smartmap map bnd in + let e' = subst_rawexpr subst e in + if bnd' == bnd && e' == e then t else CTacLet (loc, isrec, bnd', e') +| CTacTup (loc, el) -> + let el' = List.smartmap (fun e -> subst_rawexpr subst e) el in + if el' == el then t else CTacTup (loc, el') +| CTacArr (loc, el) -> + let el' = List.smartmap (fun e -> subst_rawexpr subst e) el in + if el' == el then t else CTacArr (loc, el') +| CTacLst (loc, el) -> + let el' = List.smartmap (fun e -> subst_rawexpr subst e) el in + if el' == el then t else CTacLst (loc, el') +| CTacCnv (loc, e, c) -> + let e' = subst_rawexpr subst e in + let c' = subst_rawtype subst c in + if c' == c && e' == e then t else CTacCnv (loc, e', c') +| CTacSeq (loc, e1, e2) -> + let e1' = subst_rawexpr subst e1 in + let e2' = subst_rawexpr subst e2 in + if e1' == e1 && e2' == e2 then t else CTacSeq (loc, e1', e2') +| CTacCse (loc, e, bl) -> + let map (p, e as x) = + let p' = subst_rawpattern subst p in + let e' = subst_rawexpr subst e in + if p' == p && e' == e then x else (p', e') + in + let e' = subst_rawexpr subst e in + let bl' = List.smartmap map bl in + if e' == e && bl' == bl then t else CTacCse (loc, e', bl') +| CTacRec (loc, el) -> + let map (prj, e as p) = + let prj' = subst_projection subst prj in + let e' = subst_rawexpr subst e in + if prj' == prj && e' == e then p else (prj', e') + in + let el' = List.smartmap map el in + if el' == el then t else CTacRec (loc, el') +| CTacPrj (loc, e, prj) -> + let prj' = subst_projection subst prj in + let e' = subst_rawexpr subst e in + if prj' == prj && e' == e then t else CTacPrj (loc, e', prj') +| CTacSet (loc, e, prj, r) -> + let prj' = subst_projection subst prj in + let e' = subst_rawexpr subst e in + let r' = subst_rawexpr subst r in + if prj' == prj && e' == e && r' == r then t else CTacSet (loc, e', prj', r') +| CTacExt _ -> assert false (** Should not be generated by gloabalization *) + +(** Registering *) + +let () = + let open Genintern in + let intern ist tac = + let env = match Genintern.Store.get ist.extra ltac2_env with + | None -> empty_env () + | Some env -> env + in + let loc = loc_of_tacexpr tac in + let (tac, t) = intern_rec env tac in + let () = check_elt_unit loc env t in + (ist, tac) + in + Genintern.register_intern0 wit_ltac2 intern +let () = Genintern.register_subst0 wit_ltac2 subst_expr diff --git a/src/tac2intern.mli b/src/tac2intern.mli new file mode 100644 index 0000000000..3d400a5cdd --- /dev/null +++ b/src/tac2intern.mli @@ -0,0 +1,41 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Loc.t + +val intern : raw_tacexpr -> glb_tacexpr * type_scheme +val intern_typedef : (KerName.t * int) Id.Map.t -> raw_quant_typedef -> glb_quant_typedef +val intern_open_type : raw_typexpr -> type_scheme + +(** Check that a term is a value. Only values are safe to marshall between + processes. *) +val is_value : glb_tacexpr -> bool +val check_unit : ?loc:Loc.t -> int glb_typexpr -> unit + +val subst_type : substitution -> 'a glb_typexpr -> 'a glb_typexpr +val subst_expr : substitution -> glb_tacexpr -> glb_tacexpr +val subst_quant_typedef : substitution -> glb_quant_typedef -> glb_quant_typedef +val subst_type_scheme : substitution -> type_scheme -> type_scheme + +val subst_rawexpr : substitution -> raw_tacexpr -> raw_tacexpr + +(** {5 Notations} *) + +val globalize : Id.Set.t -> raw_tacexpr -> raw_tacexpr +(** Replaces all qualified identifiers by their corresponding kernel name. The + set represents bound variables in the context. *) + +(** Errors *) + +val error_nargs_mismatch : Loc.t -> int -> int -> 'a +val error_nparams_mismatch : Loc.t -> int -> int -> 'a diff --git a/src/tac2interp.ml b/src/tac2interp.ml new file mode 100644 index 0000000000..664b7de3d6 --- /dev/null +++ b/src/tac2interp.ml @@ -0,0 +1,160 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* + let c = Tac2print.pr_constructor kn in + hov 0 (str "Uncaught Ltac2 exception:" ++ spc () ++ hov 0 c) +| _ -> raise Unhandled +end + +let val_exn = Geninterp.Val.create "ltac2:exn" + +type environment = valexpr Id.Map.t + +let empty_environment = Id.Map.empty + +let push_name ist id v = match id with +| Anonymous -> ist +| Name id -> Id.Map.add id v ist + +let get_var ist id = + try Id.Map.find id ist with Not_found -> + anomaly (str "Unbound variable " ++ Id.print id) + +let get_ref ist kn = + try pi2 (Tac2env.interp_global kn) with Not_found -> + anomaly (str "Unbound reference" ++ KerName.print kn) + +let return = Proofview.tclUNIT + +let rec interp ist = function +| GTacAtm (AtmInt n) -> return (ValInt n) +| GTacAtm (AtmStr s) -> return (ValStr (Bytes.of_string s)) +| GTacVar id -> return (get_var ist id) +| GTacRef qid -> return (get_ref ist qid) +| GTacFun (ids, e) -> + let cls = { clos_env = ist; clos_var = ids; clos_exp = e } in + return (ValCls cls) +| GTacApp (f, args) -> + interp ist f >>= fun f -> + Proofview.Monad.List.map (fun e -> interp ist e) args >>= fun args -> + interp_app f args +| GTacLet (false, el, e) -> + let fold accu (na, e) = + interp ist e >>= fun e -> + return (push_name accu na e) + in + Proofview.Monad.List.fold_left fold ist el >>= fun ist -> + interp ist e +| GTacLet (true, el, e) -> + let map (na, e) = match e with + | GTacFun (ids, e) -> + let cls = { clos_env = ist; clos_var = ids; clos_exp = e } in + na, cls + | _ -> anomaly (str "Ill-formed recursive function") + in + let fixs = List.map map el in + let fold accu (na, cls) = match na with + | Anonymous -> accu + | Name id -> Id.Map.add id (ValCls cls) accu + in + let ist = List.fold_left fold ist fixs in + (** Hack to make a cycle imperatively in the environment *) + let iter (_, e) = e.clos_env <- ist in + let () = List.iter iter fixs in + interp ist e +| GTacArr el -> + Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> + return (ValBlk (0, Array.of_list el)) +| GTacCst (_, n, []) -> return (ValInt n) +| GTacCst (_, n, el) -> + Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> + return (ValBlk (n, Array.of_list el)) +| GTacCse (e, _, cse0, cse1) -> + interp ist e >>= fun e -> interp_case ist e cse0 cse1 +| GTacWth { opn_match = e; opn_branch = cse; opn_default = def } -> + interp ist e >>= fun e -> interp_with ist e cse def +| GTacPrj (_, e, p) -> + interp ist e >>= fun e -> interp_proj ist e p +| GTacSet (_, e, p, r) -> + interp ist e >>= fun e -> + interp ist r >>= fun r -> + interp_set ist e p r +| GTacOpn (kn, el) -> + Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> + return (ValOpn (kn, Array.of_list el)) +| GTacPrm (ml, el) -> + Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> + Tac2env.interp_primitive ml el +| GTacExt e -> + let GenArg (Glbwit tag, e) = e in + let tpe = Tac2env.interp_ml_object tag in + tpe.Tac2env.ml_interp ist e >>= fun e -> return (ValExt e) + +and interp_app f args = match f with +| ValCls { clos_env = ist; clos_var = ids; clos_exp = e } -> + let rec push ist ids args = match ids, args with + | [], [] -> interp ist e + | [], _ :: _ -> interp ist e >>= fun f -> interp_app f args + | _ :: _, [] -> + let cls = { clos_env = ist; clos_var = ids; clos_exp = e } in + return (ValCls cls) + | id :: ids, arg :: args -> push (push_name ist id arg) ids args + in + push ist ids args +| ValExt _ | ValInt _ | ValBlk _ | ValStr _ | ValOpn _ -> + anomaly (str "Unexpected value shape") + +and interp_case ist e cse0 cse1 = match e with +| ValInt n -> interp ist cse0.(n) +| ValBlk (n, args) -> + let (ids, e) = cse1.(n) in + let ist = CArray.fold_left2 push_name ist ids args in + interp ist e +| ValExt _ | ValStr _ | ValCls _ | ValOpn _ -> + anomaly (str "Unexpected value shape") + +and interp_with ist e cse def = match e with +| ValOpn (kn, args) -> + let br = try Some (KNmap.find kn cse) with Not_found -> None in + begin match br with + | None -> + let (self, def) = def in + let ist = push_name ist self e in + interp ist def + | Some (self, ids, p) -> + let ist = push_name ist self e in + let ist = CArray.fold_left2 push_name ist ids args in + interp ist p + end +| ValInt _ | ValBlk _ | ValExt _ | ValStr _ | ValCls _ -> + anomaly (str "Unexpected value shape") + +and interp_proj ist e p = match e with +| ValBlk (_, args) -> + return args.(p) +| ValInt _ | ValExt _ | ValStr _ | ValCls _ | ValOpn _ -> + anomaly (str "Unexpected value shape") + +and interp_set ist e p r = match e with +| ValBlk (_, args) -> + let () = args.(p) <- r in + return (ValInt 0) +| ValInt _ | ValExt _ | ValStr _ | ValCls _ | ValOpn _ -> + anomaly (str "Unexpected value shape") diff --git a/src/tac2interp.mli b/src/tac2interp.mli new file mode 100644 index 0000000000..bf6b2d4dde --- /dev/null +++ b/src/tac2interp.mli @@ -0,0 +1,28 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* glb_tacexpr -> valexpr Proofview.tactic + +val interp_app : valexpr -> valexpr list -> valexpr Proofview.tactic + +(** {5 Exceptions} *) + +exception LtacError of KerName.t * valexpr array +(** Ltac2-defined exceptions seen from OCaml side *) + +val val_exn : Exninfo.iexn Geninterp.Val.typ +(** Toplevel representation of OCaml exceptions. Invariant: no [LtacError] + should be put into a value with tag [val_exn]. *) diff --git a/src/tac2print.ml b/src/tac2print.ml new file mode 100644 index 0000000000..e6f0582e3d --- /dev/null +++ b/src/tac2print.ml @@ -0,0 +1,296 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* str "'" ++ str (pr n) + | GTypRef (kn, []) -> pr_typref kn + | GTypRef (kn, [t]) -> + let paren = match lvl with + | T5_r | T5_l | T2 | T1 -> fun x -> x + | T0 -> paren + in + paren (pr_glbtype lvl t ++ spc () ++ pr_typref kn) + | GTypRef (kn, tl) -> + let paren = match lvl with + | T5_r | T5_l | T2 | T1 -> fun x -> x + | T0 -> paren + in + paren (str "(" ++ prlist_with_sep (fun () -> str ", ") (pr_glbtype lvl) tl ++ str ")" ++ spc () ++ pr_typref kn) + | GTypArrow (t1, t2) -> + let paren = match lvl with + | T5_r -> fun x -> x + | T5_l | T2 | T1 | T0 -> paren + in + paren (pr_glbtype T5_l t1 ++ spc () ++ str "->" ++ spc () ++ pr_glbtype T5_r t2) + | GTypTuple tl -> + let paren = match lvl with + | T5_r | T5_l -> fun x -> x + | T2 | T1 | T0 -> paren + in + paren (prlist_with_sep (fun () -> str " * ") (pr_glbtype T2) tl) + in + hov 0 (pr_glbtype lvl c) + +let pr_glbtype pr c = pr_glbtype_gen pr T5_r c + +let int_name () = + let vars = ref Int.Map.empty in + fun n -> + if Int.Map.mem n !vars then Int.Map.find n !vars + else + let num = Int.Map.cardinal !vars in + let base = num mod 26 in + let rem = num / 26 in + let name = String.make 1 (Char.chr (97 + base)) in + let suff = if Int.equal rem 0 then "" else string_of_int rem in + let name = name ^ suff in + let () = vars := Int.Map.add n name !vars in + name + +(** Term printing *) + +let pr_constructor kn = + Libnames.pr_qualid (Tac2env.shortest_qualid_of_ltac (TacConstructor kn)) + +let pr_projection kn = + Libnames.pr_qualid (Tac2env.shortest_qualid_of_projection kn) + +type exp_level = Tac2expr.exp_level = +| E5 +| E4 +| E3 +| E2 +| E1 +| E0 + +let pr_atom = function +| AtmInt n -> int n +| AtmStr s -> qstring s + +let pr_name = function +| Name id -> Id.print id +| Anonymous -> str "_" + +let find_constructor n empty def = + let rec find n = function + | [] -> assert false + | (id, []) :: rem -> + if empty then + if Int.equal n 0 then id + else find (pred n) rem + else find n rem + | (id, _ :: _) :: rem -> + if not empty then + if Int.equal n 0 then id + else find (pred n) rem + else find n rem + in + find n def + +let order_branches cbr nbr def = + let rec order cidx nidx def = match def with + | [] -> [] + | (id, []) :: rem -> + let ans = order (succ cidx) nidx rem in + (id, [], cbr.(cidx)) :: ans + | (id, _ :: _) :: rem -> + let ans = order cidx (succ nidx) rem in + let (vars, e) = nbr.(nidx) in + (id, Array.to_list vars, e) :: ans + in + order 0 0 def + +let pr_glbexpr_gen lvl c = + let rec pr_glbexpr lvl = function + | GTacAtm atm -> pr_atom atm + | GTacVar id -> Id.print id + | GTacRef gr -> + let qid = shortest_qualid_of_ltac (TacConstant gr) in + Libnames.pr_qualid qid + | GTacFun (nas, c) -> + let nas = pr_sequence pr_name nas in + let paren = match lvl with + | E0 | E1 | E2 | E3 | E4 -> paren + | E5 -> fun x -> x + in + paren (str "fun" ++ spc () ++ nas ++ spc () ++ str "=>" ++ spc () ++ + hov 0 (pr_glbexpr E5 c)) + | GTacApp (c, cl) -> + let paren = match lvl with + | E0 -> paren + | E1 | E2 | E3 | E4 | E5 -> fun x -> x + in + paren (pr_glbexpr E1 c ++ spc () ++ (pr_sequence (pr_glbexpr E0) cl)) + | GTacLet (mut, bnd, e) -> + let paren = match lvl with + | E0 | E1 | E2 | E3 | E4 -> paren + | E5 -> fun x -> x + in + let mut = if mut then str "rec" ++ spc () else mt () in + let pr_bnd (na, e) = + pr_name na ++ spc () ++ str ":=" ++ spc () ++ hov 2 (pr_glbexpr E5 e) ++ spc () + in + let bnd = prlist_with_sep (fun () -> str "with" ++ spc ()) pr_bnd bnd in + paren (str "let" ++ spc () ++ mut ++ bnd ++ str "in" ++ spc () ++ pr_glbexpr E5 e) + | GTacCst (GCaseTuple _, _, cl) -> + let paren = match lvl with + | E0 | E1 -> paren + | E2 | E3 | E4 | E5 -> fun x -> x + in + paren (prlist_with_sep (fun () -> str "," ++ spc ()) (pr_glbexpr E1) cl) + | GTacArr cl -> + mt () (** FIXME when implemented *) + | GTacCst (GCaseAlg tpe, n, cl) -> + begin match Tac2env.interp_type tpe with + | _, GTydAlg def -> + let paren = match lvl with + | E0 -> paren + | E1 | E2 | E3 | E4 | E5 -> fun x -> x + in + let id = find_constructor n (List.is_empty cl) def in + let kn = change_kn_label tpe id in + let cl = match cl with + | [] -> mt () + | _ -> spc () ++ pr_sequence (pr_glbexpr E0) cl + in + paren (pr_constructor kn ++ cl) + | _, GTydRec def -> + let args = List.combine def cl in + let pr_arg ((id, _, _), arg) = + let kn = change_kn_label tpe id in + pr_projection kn ++ spc () ++ str ":=" ++ spc () ++ pr_glbexpr E1 arg + in + let args = prlist_with_sep (fun () -> str ";" ++ spc ()) pr_arg args in + str "{" ++ spc () ++ args ++ spc () ++ str "}" + | _, (GTydDef _ | GTydOpn) -> assert false + end + | GTacCse (e, info, cst_br, ncst_br) -> + let e = pr_glbexpr E5 e in + let br = match info with + | GCaseAlg kn -> + let def = match Tac2env.interp_type kn with + | _, GTydAlg def -> def + | _, GTydDef _ | _, GTydRec _ | _, GTydOpn -> assert false + in + let br = order_branches cst_br ncst_br def in + let pr_branch (cstr, vars, p) = + let cstr = change_kn_label kn cstr in + let cstr = pr_constructor cstr in + let vars = match vars with + | [] -> mt () + | _ -> spc () ++ pr_sequence pr_name vars + in + hov 0 (str "|" ++ spc () ++ cstr ++ vars ++ spc () ++ str "=>" ++ spc () ++ + hov 2 (pr_glbexpr E5 p)) ++ spc () + in + prlist pr_branch br + | GCaseTuple n -> + let (vars, p) = ncst_br.(0) in + let p = pr_glbexpr E5 p in + let vars = prvect_with_sep (fun () -> str "," ++ spc ()) pr_name vars in + str "|" ++ spc () ++ paren vars ++ spc () ++ str "=>" ++ spc () ++ p + in + hov 0 (hov 0 (str "match" ++ spc () ++ e ++ spc () ++ str "with") ++ spc () ++ Pp.v 0 br ++ str "end") + | GTacWth wth -> + let e = pr_glbexpr E5 wth.opn_match in + let pr_pattern c self vars p = + let self = match self with + | Anonymous -> mt () + | Name id -> spc () ++ str "as" ++ spc () ++ Id.print id + in + hov 0 (str "|" ++ spc () ++ c ++ vars ++ self ++ spc () ++ str "=>" ++ spc () ++ + hov 2 (pr_glbexpr E5 p)) ++ spc () + in + let pr_branch (cstr, (self, vars, p)) = + let cstr = pr_constructor cstr in + let vars = match Array.to_list vars with + | [] -> mt () + | vars -> spc () ++ pr_sequence pr_name vars + in + pr_pattern cstr self vars p + in + let br = prlist pr_branch (KNmap.bindings wth.opn_branch) in + let (def_as, def_p) = wth.opn_default in + let def = pr_pattern (str "_") def_as (mt ()) def_p in + let br = br ++ def in + hov 0 (hov 0 (str "match" ++ spc () ++ e ++ spc () ++ str "with") ++ spc () ++ Pp.v 0 br ++ str "end") + | GTacPrj (kn, e, n) -> + let def = match Tac2env.interp_type kn with + | _, GTydRec def -> def + | _, GTydDef _ | _, GTydAlg _ | _, GTydOpn -> assert false + in + let (proj, _, _) = List.nth def n in + let proj = change_kn_label kn proj in + let proj = pr_projection proj in + let e = pr_glbexpr E0 e in + e ++ str "." ++ paren proj + | GTacSet (kn, e, n, r) -> + let def = match Tac2env.interp_type kn with + | _, GTydRec def -> def + | _, GTydDef _ | _, GTydAlg _ | _, GTydOpn -> assert false + in + let (proj, _, _) = List.nth def n in + let proj = change_kn_label kn proj in + let proj = pr_projection proj in + let e = pr_glbexpr E0 e in + let r = pr_glbexpr E1 r in + e ++ str "." ++ paren proj ++ spc () ++ str ":=" ++ spc () ++ r + | GTacOpn (kn, cl) -> + let paren = match lvl with + | E0 -> paren + | E1 | E2 | E3 | E4 | E5 -> fun x -> x + in + let c = pr_constructor kn in + paren (c ++ spc () ++ (pr_sequence (pr_glbexpr E0) cl)) + | GTacExt arg -> + let GenArg (Glbwit tag, arg) = arg in + let name = match tag with + | ExtraArg tag -> ArgT.repr tag + | _ -> assert false + in + str name ++ str ":" ++ paren (Genprint.glb_print tag arg) + | GTacPrm (prm, args) -> + let args = match args with + | [] -> mt () + | _ -> spc () ++ pr_sequence (pr_glbexpr E0) args + in + str "@external" ++ spc () ++ qstring prm.mltac_plugin ++ spc () ++ + qstring prm.mltac_tactic ++ args + in + hov 0 (pr_glbexpr lvl c) + +let pr_glbexpr c = + pr_glbexpr_gen E5 c diff --git a/src/tac2print.mli b/src/tac2print.mli new file mode 100644 index 0000000000..ddd599641d --- /dev/null +++ b/src/tac2print.mli @@ -0,0 +1,37 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* std_ppcmds +val pr_glbtype_gen : ('a -> string) -> typ_level -> 'a glb_typexpr -> std_ppcmds +val pr_glbtype : ('a -> string) -> 'a glb_typexpr -> std_ppcmds + +(** {5 Printing expressions} *) + +val pr_constructor : ltac_constructor -> std_ppcmds +val pr_projection : ltac_projection -> std_ppcmds +val pr_glbexpr_gen : exp_level -> glb_tacexpr -> std_ppcmds +val pr_glbexpr : glb_tacexpr -> std_ppcmds + +(** {5 Utilities} *) + +val int_name : unit -> (int -> string) +(** Create a function that give names to integers. The names are generated on + the fly, in the order they are encountered. *) diff --git a/tac2core.ml b/tac2core.ml deleted file mode 100644 index c82893efc2..0000000000 --- a/tac2core.ml +++ /dev/null @@ -1,646 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* t -| _ -> assert false - -let val_constr = val_tag (topwit Stdarg.wit_constr) -let val_ident = val_tag (topwit Stdarg.wit_ident) -let val_pp = Val.create "ltac2:pp" - -let extract_val (type a) (tag : a Val.typ) (Val.Dyn (tag', v)) : a = -match Val.eq tag tag' with -| None -> assert false -| Some Refl -> v - -module Core = -struct - -let t_int = coq_core "int" -let t_string = coq_core "string" -let t_array = coq_core "array" -let t_unit = coq_core "unit" -let t_list = coq_core "list" -let t_constr = coq_core "constr" -let t_ident = coq_core "ident" -let t_option = coq_core "option" - -let c_nil = coq_core "[]" -let c_cons = coq_core "::" - -let c_none = coq_core "None" -let c_some = coq_core "Some" - -end - -open Core - -let v_unit = ValInt 0 -let v_nil = ValInt 0 -let v_cons v vl = ValBlk (0, [|v; vl|]) - -module Value = -struct - -let of_unit () = v_unit - -let to_unit = function -| ValInt 0 -> () -| _ -> assert false - -let of_int n = ValInt n -let to_int = function -| ValInt n -> n -| _ -> assert false - -let of_bool b = if b then ValInt 0 else ValInt 1 - -let to_bool = function -| ValInt 0 -> true -| ValInt 1 -> false -| _ -> assert false - -let of_char n = ValInt (Char.code n) -let to_char = function -| ValInt n -> Char.chr n -| _ -> assert false - -let of_string s = ValStr s -let to_string = function -| ValStr s -> s -| _ -> assert false - -let rec of_list = function -| [] -> v_nil -| x :: l -> v_cons x (of_list l) - -let rec to_list = function -| ValInt 0 -> [] -| ValBlk (0, [|v; vl|]) -> v :: to_list vl -| _ -> assert false - -let of_ext tag c = - ValExt (Val.Dyn (tag, c)) - -let to_ext tag = function -| ValExt e -> extract_val tag e -| _ -> assert false - -let of_constr c = of_ext val_constr c -let to_constr c = to_ext val_constr c - -let of_ident c = of_ext val_ident c -let to_ident c = to_ext val_ident c - -(** FIXME: handle backtrace in Ltac2 exceptions *) -let of_exn c = match fst c with -| LtacError (kn, c) -> ValOpn (kn, c) -| _ -> of_ext val_exn c - -let to_exn c = match c with -| ValOpn (kn, c) -> (LtacError (kn, c), Exninfo.null) -| _ -> to_ext val_exn c - -let of_pp c = of_ext val_pp c -let to_pp c = to_ext val_pp c - -end - -let val_valexpr = Val.create "ltac2:valexpr" - -(** Stdlib exceptions *) - -let err_notfocussed = - LtacError (coq_core "Not_focussed", [||]) - -let err_outofbounds = - LtacError (coq_core "Out_of_bounds", [||]) - -let err_notfound = - LtacError (coq_core "Not_found", [||]) - -(** Helper functions *) - -let thaw f = interp_app f [v_unit] -let throw e = Proofview.tclLIFT (Proofview.NonLogical.raise e) - -let return x = Proofview.tclUNIT x -let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } - -let wrap f = - return () >>= fun () -> return (f ()) - -let wrap_unit f = - return () >>= fun () -> f (); return v_unit - -let pf_apply f = - Proofview.Goal.goals >>= function - | [] -> - Proofview.tclENV >>= fun env -> - Proofview.tclEVARMAP >>= fun sigma -> - f env sigma - | [gl] -> - gl >>= fun gl -> - f (Proofview.Goal.env gl) (Tacmach.New.project gl) - | _ :: _ :: _ -> - throw err_notfocussed - -(** Primitives *) - -(** Printing *) - -let prm_print : ml_tactic = function -| [pp] -> wrap_unit (fun () -> Feedback.msg_notice (Value.to_pp pp)) -| _ -> assert false - -let prm_message_of_int : ml_tactic = function -| [ValInt s] -> return (ValExt (Val.Dyn (val_pp, int s))) -| _ -> assert false - -let prm_message_of_string : ml_tactic = function -| [ValStr s] -> return (ValExt (Val.Dyn (val_pp, str (Bytes.to_string s)))) -| _ -> assert false - -let prm_message_of_constr : ml_tactic = function -| [c] -> - pf_apply begin fun env sigma -> - let c = Value.to_constr c in - let pp = Printer.pr_econstr_env env sigma c in - return (ValExt (Val.Dyn (val_pp, pp))) - end -| _ -> assert false - -let prm_message_concat : ml_tactic = function -| [m1; m2] -> - let m1 = Value.to_pp m1 in - let m2 = Value.to_pp m2 in - return (Value.of_pp (Pp.app m1 m2)) -| _ -> assert false - -(** Array *) - -let prm_array_make : ml_tactic = function -| [ValInt n; x] -> - if n < 0 || n > Sys.max_array_length then throw err_outofbounds - else wrap (fun () -> ValBlk (0, Array.make n x)) -| _ -> assert false - -let prm_array_length : ml_tactic = function -| [ValBlk (_, v)] -> return (ValInt (Array.length v)) -| _ -> assert false - -let prm_array_set : ml_tactic = function -| [ValBlk (_, v); ValInt n; x] -> - if n < 0 || n >= Array.length v then throw err_outofbounds - else wrap_unit (fun () -> v.(n) <- x) -| _ -> assert false - -let prm_array_get : ml_tactic = function -| [ValBlk (_, v); ValInt n] -> - if n < 0 || n >= Array.length v then throw err_outofbounds - else wrap (fun () -> v.(n)) -| _ -> assert false - -(** Int *) - -let prm_int_equal : ml_tactic = function -| [m; n] -> - return (Value.of_bool (Value.to_int m == Value.to_int n)) -| _ -> assert false - -let binop f : ml_tactic = function -| [m; n] -> return (Value.of_int (f (Value.to_int m) (Value.to_int n))) -| _ -> assert false - -let prm_int_compare args = binop Int.compare args -let prm_int_add args = binop (+) args -let prm_int_sub args = binop (-) args -let prm_int_mul args = binop ( * ) args - -let prm_int_neg : ml_tactic = function -| [m] -> return (Value.of_int (~- (Value.to_int m))) -| _ -> assert false - -(** String *) - -let prm_string_make : ml_tactic = function -| [n; c] -> - let n = Value.to_int n in - let c = Value.to_char c in - if n < 0 || n > Sys.max_string_length then throw err_outofbounds - else wrap (fun () -> Value.of_string (Bytes.make n c)) -| _ -> assert false - -let prm_string_length : ml_tactic = function -| [s] -> - return (Value.of_int (Bytes.length (Value.to_string s))) -| _ -> assert false - -let prm_string_set : ml_tactic = function -| [s; n; c] -> - let s = Value.to_string s in - let n = Value.to_int n in - let c = Value.to_char c in - if n < 0 || n >= Bytes.length s then throw err_outofbounds - else wrap_unit (fun () -> Bytes.set s n c) -| _ -> assert false - -let prm_string_get : ml_tactic = function -| [s; n] -> - let s = Value.to_string s in - let n = Value.to_int n in - if n < 0 || n >= Bytes.length s then throw err_outofbounds - else wrap (fun () -> Value.of_char (Bytes.get s n)) -| _ -> assert false - -(** Terms *) - -(** constr -> constr *) -let prm_constr_type : ml_tactic = function -| [c] -> - let c = Value.to_constr c in - let get_type env sigma = - Proofview.V82.wrap_exceptions begin fun () -> - let (sigma, t) = Typing.type_of env sigma c in - let t = Value.of_constr t in - Proofview.Unsafe.tclEVARS sigma <*> Proofview.tclUNIT t - end in - pf_apply get_type -| _ -> assert false - -(** constr -> constr *) -let prm_constr_equal : ml_tactic = function -| [c1; c2] -> - let c1 = Value.to_constr c1 in - let c2 = Value.to_constr c2 in - Proofview.tclEVARMAP >>= fun sigma -> - let b = EConstr.eq_constr sigma c1 c2 in - Proofview.tclUNIT (Value.of_bool b) -| _ -> assert false - -(** Error *) - -let prm_throw : ml_tactic = function -| [e] -> - let (e, info) = Value.to_exn e in - Proofview.tclLIFT (Proofview.NonLogical.raise ~info e) -| _ -> assert false - -(** Control *) - -(** exn -> 'a *) -let prm_zero : ml_tactic = function -| [e] -> - let (e, info) = Value.to_exn e in - Proofview.tclZERO ~info e -| _ -> assert false - -(** exn -> 'a *) -let prm_plus : ml_tactic = function -| [x; k] -> - Proofview.tclOR (thaw x) (fun e -> interp_app k [Value.of_exn e]) -| _ -> assert false - -(** (unit -> 'a) -> 'a *) -let prm_once : ml_tactic = function -| [f] -> Proofview.tclONCE (thaw f) -| _ -> assert false - -(** (unit -> unit) list -> unit *) -let prm_dispatch : ml_tactic = function -| [l] -> - let l = Value.to_list l in - let l = List.map (fun f -> Proofview.tclIGNORE (thaw f)) l in - Proofview.tclDISPATCH l >>= fun () -> return v_unit -| _ -> assert false - -(** (unit -> unit) list -> (unit -> unit) -> (unit -> unit) list -> unit *) -let prm_extend : ml_tactic = function -| [lft; tac; rgt] -> - let lft = Value.to_list lft in - let lft = List.map (fun f -> Proofview.tclIGNORE (thaw f)) lft in - let tac = Proofview.tclIGNORE (thaw tac) in - let rgt = Value.to_list rgt in - let rgt = List.map (fun f -> Proofview.tclIGNORE (thaw f)) rgt in - Proofview.tclEXTEND lft tac rgt >>= fun () -> return v_unit -| _ -> assert false - -(** (unit -> unit) -> unit *) -let prm_enter : ml_tactic = function -| [f] -> - let f = Proofview.tclIGNORE (thaw f) in - Proofview.tclINDEPENDENT f >>= fun () -> return v_unit -| _ -> assert false - -(** int -> int -> (unit -> 'a) -> 'a *) -let prm_focus : ml_tactic = function -| [i; j; tac] -> - let i = Value.to_int i in - let j = Value.to_int j in - Proofview.tclFOCUS i j (thaw tac) -| _ -> assert false - -(** unit -> unit *) -let prm_shelve : ml_tactic = function -| [_] -> Proofview.shelve >>= fun () -> return v_unit -| _ -> assert false - -(** unit -> unit *) -let prm_shelve_unifiable : ml_tactic = function -| [_] -> Proofview.shelve_unifiable >>= fun () -> return v_unit -| _ -> assert false - -let prm_new_goal : ml_tactic = function -| [ev] -> - let ev = Evar.unsafe_of_int (Value.to_int ev) in - Proofview.tclEVARMAP >>= fun sigma -> - if Evd.mem sigma ev then - Proofview.Unsafe.tclNEWGOALS [ev] <*> Proofview.tclUNIT v_unit - else throw err_notfound -| _ -> assert false - -(** unit -> constr *) -let prm_goal : ml_tactic = function -| [_] -> - Proofview.Goal.enter_one { enter = fun gl -> - let concl = Tacmach.New.pf_nf_concl gl in - return (Value.of_constr concl) - } -| _ -> assert false - -(** ident -> constr *) -let prm_hyp : ml_tactic = function -| [id] -> - let id = Value.to_ident id in - pf_apply begin fun env _ -> - let mem = try ignore (Environ.lookup_named id env); true with Not_found -> false in - if mem then return (Value.of_constr (EConstr.mkVar id)) - else Tacticals.New.tclZEROMSG - (str "Hypothesis " ++ quote (Id.print id) ++ str " not found") (** FIXME: Do something more sensible *) - end -| _ -> assert false - -(** (unit -> constr) -> unit *) -let prm_refine : ml_tactic = function -| [c] -> - let c = thaw c >>= fun c -> Proofview.tclUNIT ((), Value.to_constr c) in - Proofview.Goal.nf_enter { enter = fun gl -> - Refine.generic_refine ~unsafe:false c gl - } >>= fun () -> return v_unit -| _ -> assert false - - -(** Registering *) - -let () = Tac2env.define_primitive (pname "print") prm_print -let () = Tac2env.define_primitive (pname "message_of_string") prm_message_of_string -let () = Tac2env.define_primitive (pname "message_of_int") prm_message_of_int -let () = Tac2env.define_primitive (pname "message_of_constr") prm_message_of_constr -let () = Tac2env.define_primitive (pname "message_concat") prm_message_concat - -let () = Tac2env.define_primitive (pname "array_make") prm_array_make -let () = Tac2env.define_primitive (pname "array_length") prm_array_length -let () = Tac2env.define_primitive (pname "array_get") prm_array_get -let () = Tac2env.define_primitive (pname "array_set") prm_array_set - -let () = Tac2env.define_primitive (pname "string_make") prm_string_make -let () = Tac2env.define_primitive (pname "string_length") prm_string_length -let () = Tac2env.define_primitive (pname "string_get") prm_string_get -let () = Tac2env.define_primitive (pname "string_set") prm_string_set - -let () = Tac2env.define_primitive (pname "constr_type") prm_constr_type -let () = Tac2env.define_primitive (pname "constr_equal") prm_constr_equal - -let () = Tac2env.define_primitive (pname "int_equal") prm_int_equal -let () = Tac2env.define_primitive (pname "int_compare") prm_int_compare -let () = Tac2env.define_primitive (pname "int_neg") prm_int_neg -let () = Tac2env.define_primitive (pname "int_add") prm_int_add -let () = Tac2env.define_primitive (pname "int_sub") prm_int_sub -let () = Tac2env.define_primitive (pname "int_mul") prm_int_mul - -let () = Tac2env.define_primitive (pname "throw") prm_throw - -let () = Tac2env.define_primitive (pname "zero") prm_zero -let () = Tac2env.define_primitive (pname "plus") prm_plus -let () = Tac2env.define_primitive (pname "once") prm_once -let () = Tac2env.define_primitive (pname "dispatch") prm_dispatch -let () = Tac2env.define_primitive (pname "extend") prm_extend -let () = Tac2env.define_primitive (pname "enter") prm_enter - -let () = Tac2env.define_primitive (pname "focus") prm_focus -let () = Tac2env.define_primitive (pname "shelve") prm_shelve -let () = Tac2env.define_primitive (pname "shelve_unifiable") prm_shelve_unifiable -let () = Tac2env.define_primitive (pname "new_goal") prm_new_goal -let () = Tac2env.define_primitive (pname "goal") prm_goal -let () = Tac2env.define_primitive (pname "hyp") prm_hyp -let () = Tac2env.define_primitive (pname "refine") prm_refine - -(** ML types *) - -let constr_flags () = - let open Pretyping in - { - use_typeclasses = true; - solve_unification_constraints = true; - use_hook = Pfedit.solve_by_implicit_tactic (); - fail_evar = true; - expand_evars = true - } - -let open_constr_no_classes_flags () = - let open Pretyping in - { - use_typeclasses = false; - solve_unification_constraints = true; - use_hook = Pfedit.solve_by_implicit_tactic (); - fail_evar = false; - expand_evars = true - } - -(** Embed all Ltac2 data into Values *) -let to_lvar ist = - let open Pretyping in - let map e = Val.Dyn (val_valexpr, e) in - let lfun = Id.Map.map map ist in - { empty_lvar with ltac_genargs = lfun } - -let interp_constr flags ist (c, _) = - let open Pretyping in - pf_apply begin fun env sigma -> - Proofview.V82.wrap_exceptions begin fun () -> - let ist = to_lvar ist in - let (sigma, c) = understand_ltac flags env sigma ist WithoutTypeConstraint c in - let c = Val.Dyn (val_constr, c) in - Proofview.Unsafe.tclEVARS sigma >>= fun () -> - Proofview.tclUNIT c - end - end - -let () = - let interp ist c = interp_constr (constr_flags ()) ist c in - let obj = { - ml_type = t_constr; - ml_interp = interp; - } in - define_ml_object Stdarg.wit_constr obj - -let () = - let interp ist c = interp_constr (open_constr_no_classes_flags ()) ist c in - let obj = { - ml_type = t_constr; - ml_interp = interp; - } in - define_ml_object Stdarg.wit_open_constr obj - -let () = - let interp _ id = return (Val.Dyn (val_ident, id)) in - let obj = { - ml_type = t_ident; - ml_interp = interp; - } in - define_ml_object Stdarg.wit_ident obj - -let () = - let interp ist env sigma concl tac = - let fold id (Val.Dyn (tag, v)) (accu : environment) : environment = - match Val.eq tag val_valexpr with - | None -> accu - | Some Refl -> Id.Map.add id v accu - in - let ist = Id.Map.fold fold ist Id.Map.empty in - let tac = Proofview.tclIGNORE (interp ist tac) in - let c, sigma = Pfedit.refine_by_tactic env sigma concl tac in - (EConstr.of_constr c, sigma) - in - Pretyping.register_constr_interp0 wit_ltac2 interp - -(** Built-in notation scopes *) - -let add_scope s f = - Tac2entries.register_scope (Id.of_string s) f - -let scope_fail () = CErrors.user_err (str "Invalid parsing token") - -let rthunk e = - let loc = Tac2intern.loc_of_tacexpr e in - let var = [(loc, Anonymous), Some (CTypRef (loc, AbsKn Core.t_unit, []))] in - CTacFun (loc, var, e) - -let add_generic_scope s entry arg = - let parse = function - | [] -> - let scope = Extend.Aentry entry in - let act x = rthunk (CTacExt (Loc.ghost, in_gen (rawwit arg) x)) in - Tac2entries.ScopeRule (scope, act) - | _ -> scope_fail () - in - add_scope s parse - -let () = add_scope "list0" begin function -| [tok] -> - let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in - let scope = Extend.Alist0 scope in - let act l = - let l = List.map act l in - CTacLst (Loc.ghost, l) - in - Tac2entries.ScopeRule (scope, act) -| [tok; SexprStr (_, str)] -> - let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in - let sep = Extend.Atoken (CLexer.terminal str) in - let scope = Extend.Alist0sep (scope, sep) in - let act l = - let l = List.map act l in - CTacLst (Loc.ghost, l) - in - Tac2entries.ScopeRule (scope, act) -| _ -> scope_fail () -end - -let () = add_scope "list1" begin function -| [tok] -> - let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in - let scope = Extend.Alist1 scope in - let act l = - let l = List.map act l in - CTacLst (Loc.ghost, l) - in - Tac2entries.ScopeRule (scope, act) -| [tok; SexprStr (_, str)] -> - let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in - let sep = Extend.Atoken (CLexer.terminal str) in - let scope = Extend.Alist1sep (scope, sep) in - let act l = - let l = List.map act l in - CTacLst (Loc.ghost, l) - in - Tac2entries.ScopeRule (scope, act) -| _ -> scope_fail () -end - -let () = add_scope "opt" begin function -| [tok] -> - let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in - let scope = Extend.Aopt scope in - let act opt = match opt with - | None -> - CTacRef (AbsKn (TacConstructor Core.c_none)) - | Some x -> - CTacApp (Loc.ghost, CTacRef (AbsKn (TacConstructor Core.c_some)), [act x]) - in - Tac2entries.ScopeRule (scope, act) -| _ -> scope_fail () -end - -let () = add_scope "self" begin function -| [] -> - let scope = Extend.Aself in - let act tac = rthunk tac in - Tac2entries.ScopeRule (scope, act) -| _ -> scope_fail () -end - -let () = add_scope "next" begin function -| [] -> - let scope = Extend.Anext in - let act tac = rthunk tac in - Tac2entries.ScopeRule (scope, act) -| _ -> scope_fail () -end - -let () = add_scope "tactic" begin function -| [] -> - (** Default to level 5 parsing *) - let scope = Extend.Aentryl (Tac2entries.Pltac.tac2expr, 5) in - let act tac = rthunk tac in - Tac2entries.ScopeRule (scope, act) -| [SexprInt (loc, n)] -> - let () = if n < 0 || n > 5 then scope_fail () in - let scope = Extend.Aentryl (Tac2entries.Pltac.tac2expr, n) in - let act tac = rthunk tac in - Tac2entries.ScopeRule (scope, act) -| _ -> scope_fail () -end - -let () = add_generic_scope "ident" Pcoq.Prim.ident Stdarg.wit_ident -let () = add_generic_scope "constr" Pcoq.Constr.constr Stdarg.wit_constr diff --git a/tac2core.mli b/tac2core.mli deleted file mode 100644 index fc90499ac6..0000000000 --- a/tac2core.mli +++ /dev/null @@ -1,62 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* valexpr -val to_unit : valexpr -> unit - -val of_int : int -> valexpr -val to_int : valexpr -> int - -val of_bool : bool -> valexpr -val to_bool : valexpr -> bool - -val of_char : char -> valexpr -val to_char : valexpr -> char - -val of_list : valexpr list -> valexpr -val to_list : valexpr -> valexpr list - -val of_constr : EConstr.t -> valexpr -val to_constr : valexpr -> EConstr.t - -val of_exn : Exninfo.iexn -> valexpr -val to_exn : valexpr -> Exninfo.iexn - -val of_ident : Id.t -> valexpr -val to_ident : valexpr -> Id.t - -end diff --git a/tac2entries.ml b/tac2entries.ml deleted file mode 100644 index 3959e705ed..0000000000 --- a/tac2entries.ml +++ /dev/null @@ -1,645 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* obj = - declare_object {(default_object "TAC2-DEFINITION") with - cache_function = cache_tacdef; - load_function = load_tacdef; - open_function = open_tacdef; - subst_function = subst_tacdef; - classify_function = classify_tacdef} - -(** Type definition *) - -type typdef = { - typdef_local : bool; - typdef_expr : glb_quant_typedef; -} - -let change_kn_label kn id = - let (mp, dp, _) = KerName.repr kn in - KerName.make mp dp (Label.of_id id) - -let change_sp_label sp id = - let (dp, _) = Libnames.repr_path sp in - Libnames.make_path dp id - -let push_typedef visibility sp kn (_, def) = match def with -| GTydDef _ -> - Tac2env.push_type visibility sp kn -| GTydAlg cstrs -> - (** Register constructors *) - let iter (c, _) = - let spc = change_sp_label sp c in - let knc = change_kn_label kn c in - Tac2env.push_ltac visibility spc (TacConstructor knc) - in - Tac2env.push_type visibility sp kn; - List.iter iter cstrs -| GTydRec fields -> - (** Register fields *) - let iter (c, _, _) = - let spc = change_sp_label sp c in - let knc = change_kn_label kn c in - Tac2env.push_projection visibility spc knc - in - Tac2env.push_type visibility sp kn; - List.iter iter fields -| GTydOpn -> - Tac2env.push_type visibility sp kn - -let next i = - let ans = !i in - let () = incr i in - ans - -let define_typedef kn (params, def as qdef) = match def with -| GTydDef _ -> - Tac2env.define_type kn qdef -| GTydAlg cstrs -> - (** Define constructors *) - let constant = ref 0 in - let nonconstant = ref 0 in - let iter (c, args) = - let knc = change_kn_label kn c in - let tag = if List.is_empty args then next constant else next nonconstant in - let data = { - Tac2env.cdata_prms = params; - cdata_type = kn; - cdata_args = args; - cdata_indx = Some tag; - } in - Tac2env.define_constructor knc data - in - Tac2env.define_type kn qdef; - List.iter iter cstrs -| GTydRec fs -> - (** Define projections *) - let iter i (id, mut, t) = - let knp = change_kn_label kn id in - let proj = { - Tac2env.pdata_prms = params; - pdata_type = kn; - pdata_ptyp = t; - pdata_mutb = mut; - pdata_indx = i; - } in - Tac2env.define_projection knp proj - in - Tac2env.define_type kn qdef; - List.iteri iter fs -| GTydOpn -> - Tac2env.define_type kn qdef - -let perform_typdef vs ((sp, kn), def) = - let () = if not def.typdef_local then push_typedef vs sp kn def.typdef_expr in - define_typedef kn def.typdef_expr - -let load_typdef i obj = perform_typdef (Until i) obj -let open_typdef i obj = perform_typdef (Exactly i) obj - -let cache_typdef ((sp, kn), def) = - let () = push_typedef (Until 1) sp kn def.typdef_expr in - define_typedef kn def.typdef_expr - -let subst_typdef (subst, def) = - let expr' = subst_quant_typedef subst def.typdef_expr in - if expr' == def.typdef_expr then def else { def with typdef_expr = expr' } - -let classify_typdef o = Substitute o - -let inTypDef : typdef -> obj = - declare_object {(default_object "TAC2-TYPE-DEFINITION") with - cache_function = cache_typdef; - load_function = load_typdef; - open_function = open_typdef; - subst_function = subst_typdef; - classify_function = classify_typdef} - -(** Type extension *) - -type extension_data = { - edata_name : Id.t; - edata_args : int glb_typexpr list; -} - -type typext = { - typext_local : bool; - typext_prms : int; - typext_type : type_constant; - typext_expr : extension_data list; -} - -let push_typext vis sp kn def = - let iter data = - let spc = change_sp_label sp data.edata_name in - let knc = change_kn_label kn data.edata_name in - Tac2env.push_ltac vis spc (TacConstructor knc) - in - List.iter iter def.typext_expr - -let define_typext kn def = - let iter data = - let knc = change_kn_label kn data.edata_name in - let cdata = { - Tac2env.cdata_prms = def.typext_prms; - cdata_type = def.typext_type; - cdata_args = data.edata_args; - cdata_indx = None; - } in - Tac2env.define_constructor knc cdata - in - List.iter iter def.typext_expr - -let cache_typext ((sp, kn), def) = - let () = define_typext kn def in - push_typext (Until 1) sp kn def - -let perform_typext vs ((sp, kn), def) = - let () = if not def.typext_local then push_typext vs sp kn def in - define_typext kn def - -let load_typext i obj = perform_typext (Until i) obj -let open_typext i obj = perform_typext (Exactly i) obj - -let subst_typext (subst, e) = - let open Mod_subst in - let subst_data data = - let edata_args = List.smartmap (fun e -> subst_type subst e) data.edata_args in - if edata_args == data.edata_args then data - else { data with edata_args } - in - let typext_type = subst_kn subst e.typext_type in - let typext_expr = List.smartmap subst_data e.typext_expr in - if typext_type == e.typext_type && typext_expr == e.typext_expr then - e - else - { e with typext_type; typext_expr } - -let classify_typext o = Substitute o - -let inTypExt : typext -> obj = - declare_object {(default_object "TAC2-TYPE-EXTENSION") with - cache_function = cache_typext; - load_function = load_typext; - open_function = open_typext; - subst_function = subst_typext; - classify_function = classify_typext} - -(** Toplevel entries *) - -let register_ltac ?(local = false) isrec tactics = - if isrec then - let map (na, e) = (na, None, e) in - let bindings = List.map map tactics in - let map ((loc, na), e) = match na with - | Anonymous -> None - | Name id -> - let qid = Libnames.qualid_of_ident id in - let e = CTacLet (Loc.ghost, true, bindings, CTacRef (RelId (loc, qid))) in - let (e, t) = intern e in - let e = match e with - | GTacLet (true, _, e) -> assert false - | _ -> assert false - in - Some (e, t) - in - let tactics = List.map map tactics in - assert false (** FIXME *) - else - let map ((loc, na), e) = - let (e, t) = intern e in - let () = - if not (is_value e) then - user_err ~loc (str "Tactic definition must be a syntactical value") - in - let id = match na with - | Anonymous -> - user_err ~loc (str "Tactic definition must have a name") - | Name id -> id - in - let kn = Lib.make_kn id in - let exists = - try let _ = Tac2env.interp_global kn in true with Not_found -> false - in - let () = - if exists then - user_err ~loc (str "Tactic " ++ Nameops.pr_id id ++ str " already exists") - in - (id, e, t) - in - let defs = List.map map tactics in - let iter (id, e, t) = - let def = { - tacdef_local = local; - tacdef_expr = e; - tacdef_type = t; - } in - ignore (Lib.add_leaf id (inTacDef def)) - in - List.iter iter defs - -let qualid_to_ident (loc, qid) = - let (dp, id) = Libnames.repr_qualid qid in - if DirPath.is_empty dp then (loc, id) - else user_err ~loc (str "Identifier expected") - -let register_typedef ?(local = false) isrec types = - let same_name ((_, id1), _) ((_, id2), _) = Id.equal id1 id2 in - let () = match List.duplicates same_name types with - | [] -> () - | ((loc, id), _) :: _ -> - user_err ~loc (str "Multiple definition of the type name " ++ Id.print id) - in - let check ((loc, id), (params, def)) = - let same_name (_, id1) (_, id2) = Id.equal id1 id2 in - let () = match List.duplicates same_name params with - | [] -> () - | (loc, id) :: _ -> - user_err ~loc (str "The type parameter " ++ Id.print id ++ - str " occurs several times") - in - match def with - | CTydDef _ -> - if isrec then - user_err ~loc (str "The type abbreviation " ++ Id.print id ++ - str " cannot be recursive") - | CTydAlg cs -> - let same_name (id1, _) (id2, _) = Id.equal id1 id2 in - let () = match List.duplicates same_name cs with - | [] -> () - | (id, _) :: _ -> - user_err (str "Multiple definitions of the constructor " ++ Id.print id) - in - () - | CTydRec ps -> - let same_name (id1, _, _) (id2, _, _) = Id.equal id1 id2 in - let () = match List.duplicates same_name ps with - | [] -> () - | (id, _, _) :: _ -> - user_err (str "Multiple definitions of the projection " ++ Id.print id) - in - () - | CTydOpn -> - if isrec then - user_err ~loc (str "The open type declaration " ++ Id.print id ++ - str " cannot be recursive") - in - let () = List.iter check types in - let self = - if isrec then - let fold accu ((_, id), (params, _)) = - Id.Map.add id (Lib.make_kn id, List.length params) accu - in - List.fold_left fold Id.Map.empty types - else Id.Map.empty - in - let map ((_, id), def) = - let typdef = { - typdef_local = local; - typdef_expr = intern_typedef self def; - } in - (id, typdef) - in - let types = List.map map types in - let iter (id, def) = ignore (Lib.add_leaf id (inTypDef def)) in - List.iter iter types - -let register_primitive ?(local = false) (loc, id) t ml = - let t = intern_open_type t in - let rec count_arrow = function - | GTypArrow (_, t) -> 1 + count_arrow t - | _ -> 0 - in - let arrows = count_arrow (snd t) in - let () = if Int.equal arrows 0 then - user_err ~loc (str "External tactic must have at least one argument") in - let () = - try let _ = Tac2env.interp_primitive ml in () with Not_found -> - user_err ~loc (str "Unregistered primitive " ++ - quote (str ml.mltac_plugin) ++ spc () ++ quote (str ml.mltac_tactic)) - in - let init i = Id.of_string (Printf.sprintf "x%i" i) in - let names = List.init arrows init in - let bnd = List.map (fun id -> Name id) names in - let arg = List.map (fun id -> GTacVar id) names in - let e = GTacFun (bnd, GTacPrm (ml, arg)) in - let def = { - tacdef_local = local; - tacdef_expr = e; - tacdef_type = t; - } in - ignore (Lib.add_leaf id (inTacDef def)) - -let register_open ?(local = false) (loc, qid) (params, def) = - let kn = - try Tac2env.locate_type qid - with Not_found -> - user_err ~loc (str "Unbound type " ++ pr_qualid qid) - in - let (tparams, t) = Tac2env.interp_type kn in - let () = match t with - | GTydOpn -> () - | GTydAlg _ | GTydRec _ | GTydDef _ -> - user_err ~loc (str "Type " ++ pr_qualid qid ++ str " is not an open type") - in - let () = - if not (Int.equal (List.length params) tparams) then - Tac2intern.error_nparams_mismatch loc (List.length params) tparams - in - match def with - | CTydOpn -> () - | CTydAlg def -> - let intern_type t = - let tpe = CTydDef (Some t) in - let (_, ans) = intern_typedef Id.Map.empty (params, tpe) in - match ans with - | GTydDef (Some t) -> t - | _ -> assert false - in - let map (id, tpe) = - let tpe = List.map intern_type tpe in - { edata_name = id; edata_args = tpe } - in - let def = List.map map def in - let def = { - typext_local = local; - typext_type = kn; - typext_prms = tparams; - typext_expr = def; - } in - Lib.add_anonymous_leaf (inTypExt def) - | CTydRec _ | CTydDef _ -> - user_err ~loc (str "Extensions only accept inductive constructors") - -let register_type ?local isrec types = match types with -| [qid, true, def] -> - let (loc, _) = qid in - let () = if isrec then user_err ~loc (str "Extensions cannot be recursive") in - register_open ?local qid def -| _ -> - let map (qid, redef, def) = - let (loc, _) = qid in - let () = if redef then - user_err ~loc (str "Types can only be extended one by one") - in - (qualid_to_ident qid, def) - in - let types = List.map map types in - register_typedef ?local isrec types - -(** Parsing *) - -type 'a token = -| TacTerm of string -| TacNonTerm of Name.t * 'a - -type scope_rule = -| ScopeRule : (raw_tacexpr, 'a) Extend.symbol * ('a -> raw_tacexpr) -> scope_rule - -type scope_interpretation = sexpr list -> scope_rule - -let scope_table : scope_interpretation Id.Map.t ref = ref Id.Map.empty - -let register_scope id s = - scope_table := Id.Map.add id s !scope_table - -module ParseToken = -struct - -let loc_of_token = function -| SexprStr (loc, _) -> loc -| SexprInt (loc, _) -> loc -| SexprRec (loc, _, _) -> loc - -let parse_scope = function -| SexprRec (_, (loc, Some id), toks) -> - if Id.Map.mem id !scope_table then - Id.Map.find id !scope_table toks - else - CErrors.user_err ~loc (str "Unknown scope" ++ spc () ++ Nameops.pr_id id) -| tok -> - let loc = loc_of_token tok in - CErrors.user_err ~loc (str "Invalid parsing token") - -let parse_token = function -| SexprStr (_, s) -> TacTerm s -| SexprRec (_, (_, na), [tok]) -> - let na = match na with None -> Anonymous | Some id -> Name id in - let scope = parse_scope tok in - TacNonTerm (na, scope) -| tok -> - let loc = loc_of_token tok in - CErrors.user_err ~loc (str "Invalid parsing token") - -end - -let parse_scope = ParseToken.parse_scope - -type synext = { - synext_tok : sexpr list; - synext_exp : raw_tacexpr; - synext_lev : int option; - synext_loc : bool; -} - -type krule = -| KRule : - (raw_tacexpr, 'act, Loc.t -> raw_tacexpr) Extend.rule * - ((Loc.t -> (Name.t * raw_tacexpr) list -> raw_tacexpr) -> 'act) -> krule - -let rec get_rule (tok : scope_rule token list) : krule = match tok with -| [] -> KRule (Extend.Stop, fun k loc -> k loc []) -| TacNonTerm (na, ScopeRule (scope, inj)) :: tok -> - let KRule (rule, act) = get_rule tok in - let rule = Extend.Next (rule, scope) in - let act k e = act (fun loc acc -> k loc ((na, inj e) :: acc)) in - KRule (rule, act) -| TacTerm t :: tok -> - let KRule (rule, act) = get_rule tok in - let rule = Extend.Next (rule, Extend.Atoken (CLexer.terminal t)) in - let act k _ = act k in - KRule (rule, act) - -let perform_notation syn st = - let tok = List.rev_map ParseToken.parse_token syn.synext_tok in - let KRule (rule, act) = get_rule tok in - let mk loc args = - let map (na, e) = - let loc = loc_of_tacexpr e in - ((loc, na), None, e) - in - let bnd = List.map map args in - CTacLet (loc, false, bnd, syn.synext_exp) - in - let rule = Extend.Rule (rule, act mk) in - let lev = match syn.synext_lev with - | None -> None - | Some lev -> Some (string_of_int lev) - in - let rule = (lev, None, [rule]) in - ([Pcoq.ExtendRule (Pltac.tac2expr, None, (None, [rule]))], st) - -let ltac2_notation = - Pcoq.create_grammar_command "ltac2-notation" perform_notation - -let cache_synext (_, syn) = - Pcoq.extend_grammar_command ltac2_notation syn - -let open_synext i (_, syn) = - if Int.equal i 1 then Pcoq.extend_grammar_command ltac2_notation syn - -let subst_synext (subst, syn) = - let e = Tac2intern.subst_rawexpr subst syn.synext_exp in - if e == syn.synext_exp then syn else { syn with synext_exp = e } - -let classify_synext o = - if o.synext_loc then Dispose else Substitute o - -let inTac2Notation : synext -> obj = - declare_object {(default_object "TAC2-NOTATION") with - cache_function = cache_synext; - open_function = open_synext; - subst_function = subst_synext; - classify_function = classify_synext} - -let register_notation ?(local = false) tkn lev body = - (** Check that the tokens make sense *) - let entries = List.map ParseToken.parse_token tkn in - let fold accu tok = match tok with - | TacTerm _ -> accu - | TacNonTerm (Name id, _) -> Id.Set.add id accu - | TacNonTerm (Anonymous, _) -> accu - in - let ids = List.fold_left fold Id.Set.empty entries in - (** Globalize so that names are absolute *) - let body = Tac2intern.globalize ids body in - let ext = { - synext_tok = tkn; - synext_exp = body; - synext_lev = lev; - synext_loc = local; - } in - Lib.add_anonymous_leaf (inTac2Notation ext) - -(** Toplevel entries *) - -let register_struct ?local str = match str with -| StrVal (isrec, e) -> register_ltac ?local isrec e -| StrTyp (isrec, t) -> register_type ?local isrec t -| StrPrm (id, t, ml) -> register_primitive ?local id t ml -| StrSyn (tok, lev, e) -> register_notation ?local tok lev e - -(** Printing *) - -let print_ltac ref = - let (loc, qid) = qualid_of_reference ref in - let kn = - try Tac2env.locate_ltac qid - with Not_found -> user_err ~loc (str "Unknown tactic " ++ pr_qualid qid) - in - match kn with - | TacConstant kn -> - let (e, _, (_, t)) = Tac2env.interp_global kn in - let name = int_name () in - Feedback.msg_notice ( - hov 0 ( - hov 2 (pr_qualid qid ++ spc () ++ str ":" ++ spc () ++ pr_glbtype name t) ++ fnl () ++ - hov 2 (pr_qualid qid ++ spc () ++ str ":=" ++ spc () ++ pr_glbexpr e) - ) - ) - | TacConstructor kn -> - let _ = Tac2env.interp_constructor kn in - Feedback.msg_notice (hov 2 (str "Constructor" ++ spc () ++ str ":" ++ spc () ++ pr_qualid qid)) - -(** Calling tactics *) - -let solve default tac = - let status = Proof_global.with_current_proof begin fun etac p -> - let with_end_tac = if default then Some etac else None in - let (p, status) = Pfedit.solve SelectAll None tac ?with_end_tac p in - (* in case a strict subtree was completed, - go back to the top of the prooftree *) - let p = Proof.maximal_unfocus Vernacentries.command_focus p in - p, status - end in - if not status then Feedback.feedback Feedback.AddedAxiom - -let call ~default e = - let loc = loc_of_tacexpr e in - let (e, (_, t)) = intern e in - let () = check_unit ~loc t in - let tac = Tac2interp.interp Id.Map.empty e in - solve default (Proofview.tclIGNORE tac) - -(** Primitive algebraic types than can't be defined Coq-side *) - -let register_prim_alg name params def = - let id = Id.of_string name in - let def = List.map (fun (cstr, tpe) -> (Id.of_string_soft cstr, tpe)) def in - let def = (params, GTydAlg def) in - let def = { typdef_local = false; typdef_expr = def } in - ignore (Lib.add_leaf id (inTypDef def)) - -let coq_def n = KerName.make2 Tac2env.coq_prefix (Label.make n) - -let t_list = coq_def "list" - -let _ = Mltop.declare_cache_obj begin fun () -> - register_prim_alg "unit" 0 ["()", []]; - register_prim_alg "list" 1 [ - ("[]", []); - ("::", [GTypVar 0; GTypRef (t_list, [GTypVar 0])]); - ]; -end "ltac2_plugin" diff --git a/tac2entries.mli b/tac2entries.mli deleted file mode 100644 index 71e8150057..0000000000 --- a/tac2entries.mli +++ /dev/null @@ -1,57 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* rec_flag -> - (Name.t located * raw_tacexpr) list -> unit - -val register_type : ?local:bool -> rec_flag -> - (qualid located * redef_flag * raw_quant_typedef) list -> unit - -val register_primitive : ?local:bool -> - Id.t located -> raw_typexpr -> ml_tactic_name -> unit - -val register_struct : ?local:bool -> strexpr -> unit - -val register_notation : ?local:bool -> sexpr list -> int option -> - raw_tacexpr -> unit - -(** {5 Notations} *) - -type scope_rule = -| ScopeRule : (raw_tacexpr, 'a) Extend.symbol * ('a -> raw_tacexpr) -> scope_rule - -type scope_interpretation = sexpr list -> scope_rule - -val register_scope : Id.t -> scope_interpretation -> unit -(** Create a new scope with the provided name *) - -val parse_scope : sexpr -> scope_rule -(** Use this to interpret the subscopes for interpretation functions *) - -(** {5 Inspecting} *) - -val print_ltac : Libnames.reference -> unit - -(** {5 Eval loop} *) - -(** Evaluate a tactic expression in the current environment *) -val call : default:bool -> raw_tacexpr -> unit - -(** {5 Parsing entries} *) - -module Pltac : -sig -val tac2expr : raw_tacexpr Pcoq.Gram.entry -end diff --git a/tac2env.ml b/tac2env.ml deleted file mode 100644 index 5ccdd018ee..0000000000 --- a/tac2env.ml +++ /dev/null @@ -1,242 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* ValInt n -| GTacRef kn -> - let (e, _) = - try KNmap.find kn ltac_state.contents.ltac_tactics - with Not_found -> assert false - in - eval_pure e -| GTacFun (na, e) -> - ValCls { clos_env = Id.Map.empty; clos_var = na; clos_exp = e } -| GTacCst (_, n, []) -> ValInt n -| GTacCst (_, n, el) -> ValBlk (n, Array.map_of_list eval_pure el) -| GTacOpn (kn, el) -> ValOpn (kn, Array.map_of_list eval_pure el) -| GTacAtm (AtmStr _) | GTacArr _ | GTacLet _ | GTacVar _ | GTacSet _ -| GTacApp _ | GTacCse _ | GTacPrj _ | GTacPrm _ | GTacExt _ | GTacWth _ -> - anomaly (Pp.str "Term is not a syntactical value") - -let define_global kn e = - let state = !ltac_state in - ltac_state := { state with ltac_tactics = KNmap.add kn e state.ltac_tactics } - -let interp_global kn = - let (e, t) = KNmap.find kn ltac_state.contents.ltac_tactics in - (e, eval_pure e, t) - -let define_constructor kn t = - let state = !ltac_state in - ltac_state := { state with ltac_constructors = KNmap.add kn t state.ltac_constructors } - -let interp_constructor kn = KNmap.find kn ltac_state.contents.ltac_constructors - -let define_projection kn t = - let state = !ltac_state in - ltac_state := { state with ltac_projections = KNmap.add kn t state.ltac_projections } - -let interp_projection kn = KNmap.find kn ltac_state.contents.ltac_projections - -let define_type kn e = - let state = !ltac_state in - ltac_state := { state with ltac_types = KNmap.add kn e state.ltac_types } - -let interp_type kn = KNmap.find kn ltac_state.contents.ltac_types - -module ML = -struct - type t = ml_tactic_name - let compare n1 n2 = - let c = String.compare n1.mltac_plugin n2.mltac_plugin in - if Int.equal c 0 then String.compare n1.mltac_tactic n2.mltac_tactic - else c -end - -module MLMap = Map.Make(ML) - -let primitive_map = ref MLMap.empty - -let define_primitive name f = primitive_map := MLMap.add name f !primitive_map -let interp_primitive name = MLMap.find name !primitive_map - -(** Name management *) - -module FullPath = -struct - type t = full_path - let equal = eq_full_path - let to_string = string_of_path - let repr sp = - let dir,id = repr_path sp in - id, (DirPath.repr dir) -end - -type tacref = Tac2expr.tacref = -| TacConstant of ltac_constant -| TacConstructor of ltac_constructor - -module TacRef = -struct -type t = tacref -let equal r1 r2 = match r1, r2 with -| TacConstant c1, TacConstant c2 -> KerName.equal c1 c2 -| TacConstructor c1, TacConstructor c2 -> KerName.equal c1 c2 -| _ -> false -end - -module KnTab = Nametab.Make(FullPath)(KerName) -module RfTab = Nametab.Make(FullPath)(TacRef) - -type nametab = { - tab_ltac : RfTab.t; - tab_ltac_rev : full_path KNmap.t * full_path KNmap.t; - tab_type : KnTab.t; - tab_type_rev : full_path KNmap.t; - tab_proj : KnTab.t; - tab_proj_rev : full_path KNmap.t; -} - -let empty_nametab = { - tab_ltac = RfTab.empty; - tab_ltac_rev = (KNmap.empty, KNmap.empty); - tab_type = KnTab.empty; - tab_type_rev = KNmap.empty; - tab_proj = KnTab.empty; - tab_proj_rev = KNmap.empty; -} - -let nametab = Summary.ref empty_nametab ~name:"ltac2-nametab" - -let push_ltac vis sp kn = - let tab = !nametab in - let tab_ltac = RfTab.push vis sp kn tab.tab_ltac in - let (constant_map, constructor_map) = tab.tab_ltac_rev in - let tab_ltac_rev = match kn with - | TacConstant c -> (KNmap.add c sp constant_map, constructor_map) - | TacConstructor c -> (constant_map, KNmap.add c sp constructor_map) - in - nametab := { tab with tab_ltac; tab_ltac_rev } - -let locate_ltac qid = - let tab = !nametab in - RfTab.locate qid tab.tab_ltac - -let locate_extended_all_ltac qid = - let tab = !nametab in - RfTab.find_prefixes qid tab.tab_ltac - -let shortest_qualid_of_ltac kn = - let tab = !nametab in - let sp = match kn with - | TacConstant c -> KNmap.find c (fst tab.tab_ltac_rev) - | TacConstructor c -> KNmap.find c (snd tab.tab_ltac_rev) - in - RfTab.shortest_qualid Id.Set.empty sp tab.tab_ltac - -let push_type vis sp kn = - let tab = !nametab in - let tab_type = KnTab.push vis sp kn tab.tab_type in - let tab_type_rev = KNmap.add kn sp tab.tab_type_rev in - nametab := { tab with tab_type; tab_type_rev } - -let locate_type qid = - let tab = !nametab in - KnTab.locate qid tab.tab_type - -let locate_extended_all_type qid = - let tab = !nametab in - KnTab.find_prefixes qid tab.tab_type - -let shortest_qualid_of_type kn = - let tab = !nametab in - let sp = KNmap.find kn tab.tab_type_rev in - KnTab.shortest_qualid Id.Set.empty sp tab.tab_type - -let push_projection vis sp kn = - let tab = !nametab in - let tab_proj = KnTab.push vis sp kn tab.tab_proj in - let tab_proj_rev = KNmap.add kn sp tab.tab_proj_rev in - nametab := { tab with tab_proj; tab_proj_rev } - -let locate_projection qid = - let tab = !nametab in - KnTab.locate qid tab.tab_proj - -let locate_extended_all_projection qid = - let tab = !nametab in - KnTab.find_prefixes qid tab.tab_proj - -let shortest_qualid_of_projection kn = - let tab = !nametab in - let sp = KNmap.find kn tab.tab_proj_rev in - KnTab.shortest_qualid Id.Set.empty sp tab.tab_proj - -type 'a ml_object = { - ml_type : type_constant; - ml_interp : environment -> 'a -> Geninterp.Val.t Proofview.tactic; -} - -module MLTypeObj = -struct - type ('a, 'b, 'c) obj = 'b ml_object - let name = "ltac2_ml_type" - let default _ = None -end - -module MLType = Genarg.Register(MLTypeObj) - -let define_ml_object t tpe = MLType.register0 t tpe -let interp_ml_object t = MLType.obj t - -(** Absolute paths *) - -let coq_prefix = - MPfile (DirPath.make (List.map Id.of_string ["Init"; "ltac2"; "Coq"])) - -(** Generic arguments *) - -let wit_ltac2 = Genarg.make0 "ltac2" diff --git a/tac2env.mli b/tac2env.mli deleted file mode 100644 index c4b8c1e0ca..0000000000 --- a/tac2env.mli +++ /dev/null @@ -1,106 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* (glb_tacexpr * type_scheme) -> unit -val interp_global : ltac_constant -> (glb_tacexpr * valexpr * type_scheme) - -(** {5 Toplevel definition of types} *) - -val define_type : type_constant -> glb_quant_typedef -> unit -val interp_type : type_constant -> glb_quant_typedef - -(** {5 Toplevel definition of algebraic constructors} *) - -type constructor_data = { - cdata_prms : int; - (** Type parameters *) - cdata_type : type_constant; - (** Inductive definition to which the constructor pertains *) - cdata_args : int glb_typexpr list; - (** Types of the constructor arguments *) - cdata_indx : int option; - (** Index of the constructor in the ADT. Numbering is duplicated between - argumentless and argument-using constructors, e.g. in type ['a option] - [None] and [Some] have both index 0. This field is empty whenever the - constructor is a member of an open type. *) -} - -val define_constructor : ltac_constructor -> constructor_data -> unit -val interp_constructor : ltac_constructor -> constructor_data - -(** {5 Toplevel definition of projections} *) - -type projection_data = { - pdata_prms : int; - (** Type parameters *) - pdata_type : type_constant; - (** Record definition to which the projection pertains *) - pdata_ptyp : int glb_typexpr; - (** Type of the projection *) - pdata_mutb : bool; - (** Whether the field is mutable *) - pdata_indx : int; - (** Index of the projection *) -} - -val define_projection : ltac_projection -> projection_data -> unit -val interp_projection : ltac_projection -> projection_data - -(** {5 Name management} *) - -val push_ltac : visibility -> full_path -> tacref -> unit -val locate_ltac : qualid -> tacref -val locate_extended_all_ltac : qualid -> tacref list -val shortest_qualid_of_ltac : tacref -> qualid - -val push_type : visibility -> full_path -> type_constant -> unit -val locate_type : qualid -> type_constant -val locate_extended_all_type : qualid -> type_constant list -val shortest_qualid_of_type : type_constant -> qualid - -val push_projection : visibility -> full_path -> ltac_projection -> unit -val locate_projection : qualid -> ltac_projection -val locate_extended_all_projection : qualid -> ltac_projection list -val shortest_qualid_of_projection : ltac_projection -> qualid - -(** {5 Toplevel definitions of ML tactics} *) - -(** This state is not part of the summary, contrarily to the ones above. It is - intended to be used from ML plugins to register ML-side functions. *) - -val define_primitive : ml_tactic_name -> ml_tactic -> unit -val interp_primitive : ml_tactic_name -> ml_tactic - -(** {5 ML primitive types} *) - -type 'a ml_object = { - ml_type : type_constant; - ml_interp : environment -> 'a -> Geninterp.Val.t Proofview.tactic; -} - -val define_ml_object : ('a, 'b, 'c) genarg_type -> 'b ml_object -> unit -val interp_ml_object : ('a, 'b, 'c) genarg_type -> 'b ml_object - -(** {5 Absolute paths} *) - -val coq_prefix : ModPath.t -(** Path where primitive datatypes are defined in Ltac2 plugin. *) - -(** {5 Generic arguments} *) - -val wit_ltac2 : (raw_tacexpr, glb_tacexpr, Util.Empty.t) genarg_type diff --git a/tac2expr.mli b/tac2expr.mli deleted file mode 100644 index acdad9bab4..0000000000 --- a/tac2expr.mli +++ /dev/null @@ -1,195 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* valexpr Proofview.tactic - -type environment = valexpr Id.Map.t diff --git a/tac2intern.ml b/tac2intern.ml deleted file mode 100644 index 756bbe3076..0000000000 --- a/tac2intern.ml +++ /dev/null @@ -1,1452 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* elt -> bool -val create : unit -> 'a t -val fresh : 'a t -> elt -val find : elt -> 'a t -> (elt * 'a option) -val union : elt -> elt -> 'a t -> unit -val set : elt -> 'a -> 'a t -> unit -module Map : -sig - type key = elt - type +'a t - val empty : 'a t - val add : key -> 'a -> 'a t -> 'a t - val mem : key -> 'a t -> bool - val find : key -> 'a t -> 'a - val exists : (key -> 'a -> bool) -> 'a t -> bool -end -end -= -struct -type elt = int -let equal = Int.equal -module Map = Int.Map - -type 'a node = -| Canon of int * 'a option -| Equiv of elt - -type 'a t = { - mutable uf_data : 'a node array; - mutable uf_size : int; -} - -let resize p = - if Int.equal (Array.length p.uf_data) p.uf_size then begin - let nsize = 2 * p.uf_size + 1 in - let v = Array.make nsize (Equiv 0) in - Array.blit p.uf_data 0 v 0 (Array.length p.uf_data); - p.uf_data <- v; - end - -let create () = { uf_data = [||]; uf_size = 0 } - -let fresh p = - resize p; - let n = p.uf_size in - p.uf_data.(n) <- (Canon (1, None)); - p.uf_size <- n + 1; - n - -let rec lookup n p = - let node = Array.get p.uf_data n in - match node with - | Canon (size, v) -> n, size, v - | Equiv y -> - let ((z, _, _) as res) = lookup y p in - if not (Int.equal z y) then Array.set p.uf_data n (Equiv z); - res - -let find n p = - let (x, _, v) = lookup n p in (x, v) - -let union x y p = - let ((x, size1, _) as xcan) = lookup x p in - let ((y, size2, _) as ycan) = lookup y p in - let xcan, ycan = if size1 < size2 then xcan, ycan else ycan, xcan in - let x, _, xnode = xcan in - let y, _, ynode = ycan in - assert (Option.is_empty xnode); - assert (Option.is_empty ynode); - p.uf_data.(x) <- Equiv y; - p.uf_data.(y) <- Canon (size1 + size2, None) - -let set x v p = - let (x, s, v') = lookup x p in - assert (Option.is_empty v'); - p.uf_data.(x) <- Canon (s, Some v) - -end - -type mix_var = -| GVar of UF.elt -| LVar of int - -type mix_type_scheme = int * mix_var glb_typexpr - -type environment = { - env_var : mix_type_scheme Id.Map.t; - (** Type schemes of bound variables *) - env_cst : UF.elt glb_typexpr UF.t; - (** Unification state *) - env_als : UF.elt Id.Map.t ref; - (** Map user-facing type variables to unification variables *) - env_opn : bool; - (** Accept unbound type variables *) - env_rec : (KerName.t * int) Id.Map.t; - (** Recursive type definitions *) -} - -let empty_env () = { - env_var = Id.Map.empty; - env_cst = UF.create (); - env_als = ref Id.Map.empty; - env_opn = true; - env_rec = Id.Map.empty; -} - -let env_name env = - (** Generate names according to a provided environment *) - let mk num = - let base = num mod 26 in - let rem = num / 26 in - let name = String.make 1 (Char.chr (97 + base)) in - let suff = if Int.equal rem 0 then "" else string_of_int rem in - let name = name ^ suff in - name - in - let fold id elt acc = UF.Map.add elt (Id.to_string id) acc in - let vars = Id.Map.fold fold env.env_als.contents UF.Map.empty in - let vars = ref vars in - let rec fresh n = - let name = mk n in - if UF.Map.exists (fun _ name' -> String.equal name name') !vars then fresh (succ n) - else name - in - fun n -> - if UF.Map.mem n !vars then UF.Map.find n !vars - else - let ans = fresh 0 in - let () = vars := UF.Map.add n ans !vars in - ans - -let ltac2_env : environment Genintern.Store.field = - Genintern.Store.field () - -let fresh_id env = UF.fresh env.env_cst - -let get_alias (loc, id) env = - try Id.Map.find id env.env_als.contents - with Not_found -> - if env.env_opn then - let n = fresh_id env in - let () = env.env_als := Id.Map.add id n env.env_als.contents in - n - else user_err ~loc (str "Unbound type parameter " ++ Id.print id) - -let push_name id t env = match id with -| Anonymous -> env -| Name id -> { env with env_var = Id.Map.add id t env.env_var } - -let loc_of_tacexpr = function -| CTacAtm (loc, _) -> loc -| CTacRef (RelId (loc, _)) -> loc -| CTacRef (AbsKn _) -> Loc.ghost -| CTacFun (loc, _, _) -> loc -| CTacApp (loc, _, _) -> loc -| CTacLet (loc, _, _, _) -> loc -| CTacTup (loc, _) -> loc -| CTacArr (loc, _) -> loc -| CTacLst (loc, _) -> loc -| CTacCnv (loc, _, _) -> loc -| CTacSeq (loc, _, _) -> loc -| CTacCse (loc, _, _) -> loc -| CTacRec (loc, _) -> loc -| CTacPrj (loc, _, _) -> loc -| CTacSet (loc, _, _, _) -> loc -| CTacExt (loc, _) -> loc - -let loc_of_patexpr = function -| CPatAny loc -> loc -| CPatRef (loc, _, _) -> loc -| CPatTup (loc, _) -> loc - -let error_nargs_mismatch loc nargs nfound = - user_err ~loc (str "Constructor expects " ++ int nargs ++ - str " arguments, but is applied to " ++ int nfound ++ - str " arguments") - -let error_nparams_mismatch loc nargs nfound = - user_err ~loc (str "Type expects " ++ int nargs ++ - str " arguments, but is applied to " ++ int nfound ++ - str " arguments") - -let rec subst_type subst (t : 'a glb_typexpr) = match t with -| GTypVar id -> subst id -| GTypArrow (t1, t2) -> GTypArrow (subst_type subst t1, subst_type subst t2) -| GTypTuple tl -> GTypTuple (List.map (fun t -> subst_type subst t) tl) -| GTypRef (qid, args) -> - GTypRef (qid, List.map (fun t -> subst_type subst t) args) - -let rec intern_type env (t : raw_typexpr) : UF.elt glb_typexpr = match t with -| CTypVar (loc, Name id) -> GTypVar (get_alias (loc, id) env) -| CTypVar (_, Anonymous) -> GTypVar (fresh_id env) -| CTypRef (loc, rel, args) -> - let (kn, nparams) = match rel with - | RelId (loc, qid) -> - let (dp, id) = repr_qualid qid in - if DirPath.is_empty dp && Id.Map.mem id env.env_rec then - Id.Map.find id env.env_rec - else - let kn = - try Tac2env.locate_type qid - with Not_found -> - user_err ~loc (str "Unbound type constructor " ++ pr_qualid qid) - in - let (nparams, _) = Tac2env.interp_type kn in - (kn, nparams) - | AbsKn kn -> - let (nparams, _) = Tac2env.interp_type kn in - (kn, nparams) - in - let nargs = List.length args in - let () = - if not (Int.equal nparams nargs) then - let loc, qid = match rel with - | RelId lid -> lid - | AbsKn kn -> loc, shortest_qualid_of_type kn - in - user_err ~loc (strbrk "The type constructor " ++ pr_qualid qid ++ - strbrk " expects " ++ int nparams ++ strbrk " argument(s), but is here \ - applied to " ++ int nargs ++ strbrk "argument(s)") - in - GTypRef (kn, List.map (fun t -> intern_type env t) args) -| CTypArrow (loc, t1, t2) -> GTypArrow (intern_type env t1, intern_type env t2) -| CTypTuple (loc, tl) -> GTypTuple (List.map (fun t -> intern_type env t) tl) - -let fresh_type_scheme env (t : type_scheme) : UF.elt glb_typexpr = - let (n, t) = t in - let subst = Array.init n (fun _ -> fresh_id env) in - let substf i = GTypVar subst.(i) in - subst_type substf t - -let fresh_mix_type_scheme env (t : mix_type_scheme) : UF.elt glb_typexpr = - let (n, t) = t in - let subst = Array.init n (fun _ -> fresh_id env) in - let substf = function - | LVar i -> GTypVar subst.(i) - | GVar n -> GTypVar n - in - subst_type substf t - -let fresh_reftype env (kn : KerName.t) = - let (n, _) = Tac2env.interp_type kn in - let subst = Array.init n (fun _ -> fresh_id env) in - let t = GTypRef (kn, Array.map_to_list (fun i -> GTypVar i) subst) in - (subst, t) - -(** First-order unification algorithm *) - -let is_unfoldable kn = match snd (Tac2env.interp_type kn) with -| GTydDef (Some _) -> true -| GTydDef None | GTydAlg _ | GTydRec _ | GTydOpn -> false - -let unfold env kn args = - let (nparams, def) = Tac2env.interp_type kn in - let def = match def with - | GTydDef (Some t) -> t - | _ -> assert false - in - let args = Array.of_list args in - let subst n = args.(n) in - subst_type subst def - -(** View function, allows to ensure head normal forms *) -let rec kind env t = match t with -| GTypVar id -> - let (id, v) = UF.find id env.env_cst in - begin match v with - | None -> GTypVar id - | Some t -> kind env t - end -| GTypRef (kn, tl) -> - if is_unfoldable kn then kind env (unfold env kn tl) else t -| GTypArrow _ | GTypTuple _ -> t - -exception Occur - -let rec occur_check env id t = match kind env t with -| GTypVar id' -> if UF.equal id id' then raise Occur -| GTypArrow (t1, t2) -> - let () = occur_check env id t1 in - occur_check env id t2 -| GTypTuple tl -> - List.iter (fun t -> occur_check env id t) tl -| GTypRef (kn, tl) -> - List.iter (fun t -> occur_check env id t) tl - -exception CannotUnify of UF.elt glb_typexpr * UF.elt glb_typexpr - -let unify_var env id t = match kind env t with -| GTypVar id' -> - if not (UF.equal id id') then UF.union id id' env.env_cst -| GTypArrow _ | GTypRef _ | GTypTuple _ -> - try - let () = occur_check env id t in - UF.set id t env.env_cst - with Occur -> raise (CannotUnify (GTypVar id, t)) - -let rec unify env t1 t2 = match kind env t1, kind env t2 with -| GTypVar id, t | t, GTypVar id -> - unify_var env id t -| GTypArrow (t1, u1), GTypArrow (t2, u2) -> - let () = unify env t1 t2 in - unify env u1 u2 -| GTypTuple tl1, GTypTuple tl2 -> - if Int.equal (List.length tl1) (List.length tl2) then - List.iter2 (fun t1 t2 -> unify env t1 t2) tl1 tl2 - else raise (CannotUnify (t1, t2)) -| GTypRef (kn1, tl1), GTypRef (kn2, tl2) -> - if KerName.equal kn1 kn2 then - List.iter2 (fun t1 t2 -> unify env t1 t2) tl1 tl2 - else raise (CannotUnify (t1, t2)) -| _ -> raise (CannotUnify (t1, t2)) - -let unify loc env t1 t2 = - try unify env t1 t2 - with CannotUnify (u1, u2) -> - let name = env_name env in - user_err ~loc (str "This expression has type " ++ pr_glbtype name t1 ++ - str " but an expression what expected of type " ++ pr_glbtype name t2) - -(** Term typing *) - -let is_pure_constructor kn = - match snd (Tac2env.interp_type kn) with - | GTydAlg _ | GTydOpn -> true - | GTydRec fields -> - let is_pure (_, mut, _) = not mut in - List.for_all is_pure fields - | GTydDef _ -> assert false (** Type definitions have no constructors *) - -let rec is_value = function -| GTacAtm (AtmInt _) | GTacVar _ | GTacRef _ | GTacFun _ -> true -| GTacAtm (AtmStr _) | GTacApp _ | GTacLet _ -> false -| GTacCst (GCaseTuple _, _, el) -> List.for_all is_value el -| GTacCst (_, _, []) -> true -| GTacOpn (_, el) -> List.for_all is_value el -| GTacCst (GCaseAlg kn, _, el) -> is_pure_constructor kn && List.for_all is_value el -| GTacArr _ | GTacCse _ | GTacPrj _ | GTacSet _ | GTacExt _ | GTacPrm _ -| GTacWth _ -> false - -let is_rec_rhs = function -| GTacFun _ -> true -| GTacAtm _ | GTacVar _ | GTacRef _ | GTacApp _ | GTacLet _ | GTacPrj _ -| GTacSet _ | GTacArr _ | GTacExt _ | GTacPrm _ | GTacCst _ -| GTacCse _ | GTacOpn _ | GTacWth _ -> false - -let rec fv_type f t accu = match t with -| GTypVar id -> f id accu -| GTypArrow (t1, t2) -> fv_type f t1 (fv_type f t2 accu) -| GTypTuple tl -> List.fold_left (fun accu t -> fv_type f t accu) accu tl -| GTypRef (kn, tl) -> List.fold_left (fun accu t -> fv_type f t accu) accu tl - -let fv_env env = - let rec f id accu = match UF.find id env.env_cst with - | id, None -> UF.Map.add id () accu - | _, Some t -> fv_type f t accu - in - let fold_var id (_, t) accu = - let fmix id accu = match id with - | LVar _ -> accu - | GVar id -> f id accu - in - fv_type fmix t accu - in - let fv_var = Id.Map.fold fold_var env.env_var UF.Map.empty in - let fold_als _ id accu = f id accu in - Id.Map.fold fold_als !(env.env_als) fv_var - -let abstract_var env (t : UF.elt glb_typexpr) : mix_type_scheme = - let fv = fv_env env in - let count = ref 0 in - let vars = ref UF.Map.empty in - let rec subst id = - let (id, t) = UF.find id env.env_cst in - match t with - | None -> - if UF.Map.mem id fv then GTypVar (GVar id) - else - begin try UF.Map.find id !vars - with Not_found -> - let n = !count in - let var = GTypVar (LVar n) in - let () = incr count in - let () = vars := UF.Map.add id var !vars in - var - end - | Some t -> subst_type subst t - in - let t = subst_type subst t in - (!count, t) - -let monomorphic (t : UF.elt glb_typexpr) : mix_type_scheme = - let subst id = GTypVar (GVar id) in - (0, subst_type subst t) - -let warn_not_unit = - CWarnings.create ~name:"not-unit" ~category:"ltac" - (fun () -> strbrk "The following expression should have type unit.") - -let warn_redundant_clause = - CWarnings.create ~name:"redundant-clause" ~category:"ltac" - (fun () -> strbrk "The following clause is redundant.") - -let check_elt_unit loc env t = - let maybe_unit = match kind env t with - | GTypVar _ -> true - | GTypArrow _ | GTypTuple _ -> false - | GTypRef (kn, _) -> KerName.equal kn t_unit - in - if not maybe_unit then warn_not_unit ~loc () - -let check_elt_empty loc env t = match kind env t with -| GTypVar _ -> - user_err ~loc (str "Cannot infer an empty type for this expression") -| GTypArrow _ | GTypTuple _ -> - let name = env_name env in - user_err ~loc (str "Type " ++ pr_glbtype name t ++ str " is not an empty type") -| GTypRef (kn, _) -> - let def = Tac2env.interp_type kn in - match def with - | _, GTydAlg [] -> kn - | _ -> - let name = env_name env in - user_err ~loc (str "Type " ++ pr_glbtype name t ++ str " is not an empty type") - -let check_unit ?loc t = - let maybe_unit = match t with - | GTypVar _ -> true - | GTypArrow _ | GTypTuple _ -> false - | GTypRef (kn, _) -> KerName.equal kn t_unit - in - if not maybe_unit then warn_not_unit ?loc () - -let check_redundant_clause = function -| [] -> () -| (p, _) :: _ -> warn_redundant_clause ~loc:(loc_of_patexpr p) () - -let get_variable0 mem var = match var with -| RelId (loc, qid) -> - let (dp, id) = repr_qualid qid in - if DirPath.is_empty dp && mem id then ArgVar (loc, id) - else - let kn = - try Tac2env.locate_ltac qid - with Not_found -> - CErrors.user_err ~loc (str "Unbound value " ++ pr_qualid qid) - in - ArgArg kn -| AbsKn kn -> ArgArg kn - -let get_variable env var = - let mem id = Id.Map.mem id env.env_var in - get_variable0 mem var - -let get_constructor env var = match var with -| RelId (loc, qid) -> - let c = try Some (Tac2env.locate_ltac qid) with Not_found -> None in - begin match c with - | Some (TacConstructor knc) -> - let kn = Tac2env.interp_constructor knc in - ArgArg (kn, knc) - | Some (TacConstant _) -> - CErrors.user_err ~loc (str "The term " ++ pr_qualid qid ++ - str " is not the constructor of an inductive type.") - | None -> - let (dp, id) = repr_qualid qid in - if DirPath.is_empty dp then ArgVar (loc, id) - else CErrors.user_err ~loc (str "Unbound constructor " ++ pr_qualid qid) - end -| AbsKn knc -> - let kn = Tac2env.interp_constructor knc in - ArgArg (kn, knc) - -let get_projection var = match var with -| RelId (loc, qid) -> - let kn = try Tac2env.locate_projection qid with Not_found -> - user_err ~loc (pr_qualid qid ++ str " is not a projection") - in - Tac2env.interp_projection kn -| AbsKn kn -> - Tac2env.interp_projection kn - -let intern_atm env = function -| AtmInt n -> (GTacAtm (AtmInt n), GTypRef (t_int, [])) -| AtmStr s -> (GTacAtm (AtmStr s), GTypRef (t_string, [])) - -let invalid_pattern ~loc kn t = - let pt = match t with - | GCaseAlg kn' -> pr_typref kn - | GCaseTuple n -> str "tuple" - in - user_err ~loc (str "Invalid pattern, expected a pattern for type " ++ - pr_typref kn ++ str ", found a pattern of type " ++ pt) (** FIXME *) - -(** Pattern view *) - -type glb_patexpr = -| GPatVar of Name.t -| GPatRef of ltac_constructor * glb_patexpr list -| GPatTup of glb_patexpr list - -let rec intern_patexpr env = function -| CPatAny _ -> GPatVar Anonymous -| CPatRef (_, qid, []) -> - begin match get_constructor env qid with - | ArgVar (_, id) -> GPatVar (Name id) - | ArgArg (_, kn) -> GPatRef (kn, []) - end -| CPatRef (_, qid, pl) -> - begin match get_constructor env qid with - | ArgVar (loc, id) -> - user_err ~loc (str "Unbound constructor " ++ Nameops.pr_id id) - | ArgArg (_, kn) -> GPatRef (kn, List.map (fun p -> intern_patexpr env p) pl) - end -| CPatTup (_, pl) -> - GPatTup (List.map (fun p -> intern_patexpr env p) pl) - -type pattern_kind = -| PKind_empty -| PKind_variant of type_constant -| PKind_open of type_constant -| PKind_tuple of int -| PKind_any - -let get_pattern_kind env pl = match pl with -| [] -> PKind_empty -| p :: pl -> - let rec get_kind (p, _) pl = match intern_patexpr env p with - | GPatVar _ -> - begin match pl with - | [] -> PKind_any - | p :: pl -> get_kind p pl - end - | GPatRef (kn, pl) -> - let data = Tac2env.interp_constructor kn in - if Option.is_empty data.cdata_indx then PKind_open data.cdata_type - else PKind_variant data.cdata_type - | GPatTup tp -> PKind_tuple (List.length tp) - in - get_kind p pl - -(** Internalization *) - -let is_constructor env qid = match get_variable env qid with -| ArgArg (TacConstructor _) -> true -| _ -> false - -let rec intern_rec env = function -| CTacAtm (_, atm) -> intern_atm env atm -| CTacRef qid as e -> - begin match get_variable env qid with - | ArgVar (_, id) -> - let sch = Id.Map.find id env.env_var in - (GTacVar id, fresh_mix_type_scheme env sch) - | ArgArg (TacConstant kn) -> - let (_, _, sch) = Tac2env.interp_global kn in - (GTacRef kn, fresh_type_scheme env sch) - | ArgArg (TacConstructor kn) -> - let loc = loc_of_tacexpr e in - intern_constructor env loc kn [] - end -| CTacFun (loc, bnd, e) -> - let fold (env, bnd, tl) ((_, na), t) = - let t = match t with - | None -> GTypVar (fresh_id env) - | Some t -> intern_type env t - in - let env = push_name na (monomorphic t) env in - (env, na :: bnd, t :: tl) - in - let (env, bnd, tl) = List.fold_left fold (env, [], []) bnd in - let bnd = List.rev bnd in - let (e, t) = intern_rec env e in - let t = List.fold_left (fun accu t -> GTypArrow (t, accu)) t tl in - (GTacFun (bnd, e), t) -| CTacApp (loc, CTacRef qid, args) as e when is_constructor env qid -> - let kn = match get_variable env qid with - | ArgArg (TacConstructor kn) -> kn - | _ -> assert false - in - let loc = loc_of_tacexpr e in - intern_constructor env loc kn args -| CTacApp (loc, f, args) -> - let (f, ft) = intern_rec env f in - let fold arg (args, t) = - let (arg, argt) = intern_rec env arg in - (arg :: args, GTypArrow (argt, t)) - in - let ret = GTypVar (fresh_id env) in - let (args, t) = List.fold_right fold args ([], ret) in - let () = unify loc env ft t in - (GTacApp (f, args), ret) -| CTacLet (loc, false, el, e) -> - let fold accu ((loc, na), _, _) = match na with - | Anonymous -> accu - | Name id -> - if Id.Set.mem id accu then - user_err ~loc (str "Variable " ++ Id.print id ++ str " is bound several \ - times in this matching") - else Id.Set.add id accu - in - let _ = List.fold_left fold Id.Set.empty el in - let fold ((loc, na), tc, e) (el, p) = - let (e, t) = intern_rec env e in - let () = match tc with - | None -> () - | Some tc -> - let tc = intern_type env tc in - unify loc env t tc - in - let t = if is_value e then abstract_var env t else monomorphic t in - ((na, e) :: el), ((na, t) :: p) - in - let (el, p) = List.fold_right fold el ([], []) in - let nenv = List.fold_left (fun accu (na, t) -> push_name na t env) env p in - let (e, t) = intern_rec nenv e in - (GTacLet (false, el, e), t) -| CTacLet (loc, true, el, e) -> - intern_let_rec env loc el e -| CTacTup (loc, []) -> - (GTacCst (GCaseAlg t_unit, 0, []), GTypRef (t_unit, [])) -| CTacTup (loc, el) -> - let fold e (el, tl) = - let (e, t) = intern_rec env e in - (e :: el, t :: tl) - in - let (el, tl) = List.fold_right fold el ([], []) in - (GTacCst (GCaseTuple (List.length el), 0, el), GTypTuple tl) -| CTacArr (loc, []) -> - let id = fresh_id env in - (GTacArr [], GTypRef (t_int, [GTypVar id])) -| CTacArr (loc, e0 :: el) -> - let (e0, t0) = intern_rec env e0 in - let fold e el = intern_rec_with_constraint env e t0 :: el in - let el = e0 :: List.fold_right fold el [] in - (GTacArr el, GTypRef (t_array, [t0])) -| CTacLst (loc, []) -> - let id = fresh_id env in - (c_nil, GTypRef (t_list, [GTypVar id])) -| CTacLst (loc, e0 :: el) -> - let (e0, t0) = intern_rec env e0 in - let fold e el = c_cons (intern_rec_with_constraint env e t0) el in - let el = c_cons e0 (List.fold_right fold el c_nil) in - (el, GTypRef (t_list, [t0])) -| CTacCnv (loc, e, tc) -> - let (e, t) = intern_rec env e in - let tc = intern_type env tc in - let () = unify loc env t tc in - (e, tc) -| CTacSeq (loc, e1, e2) -> - let (e1, t1) = intern_rec env e1 in - let (e2, t2) = intern_rec env e2 in - let () = check_elt_unit loc env t1 in - (GTacLet (false, [Anonymous, e1], e2), t2) -| CTacCse (loc, e, pl) -> - intern_case env loc e pl -| CTacRec (loc, fs) -> - intern_record env loc fs -| CTacPrj (loc, e, proj) -> - let pinfo = get_projection proj in - let loc = loc_of_tacexpr e in - let (e, t) = intern_rec env e in - let subst = Array.init pinfo.pdata_prms (fun _ -> fresh_id env) in - let params = Array.map_to_list (fun i -> GTypVar i) subst in - let exp = GTypRef (pinfo.pdata_type, params) in - let () = unify loc env t exp in - let substf i = GTypVar subst.(i) in - let ret = subst_type substf pinfo.pdata_ptyp in - (GTacPrj (pinfo.pdata_type, e, pinfo.pdata_indx), ret) -| CTacSet (loc, e, proj, r) -> - let pinfo = get_projection proj in - let () = - if not pinfo.pdata_mutb then - let loc = match proj with - | RelId (loc, _) -> loc - | AbsKn _ -> Loc.ghost - in - user_err ~loc (str "Field is not mutable") - in - let subst = Array.init pinfo.pdata_prms (fun _ -> fresh_id env) in - let params = Array.map_to_list (fun i -> GTypVar i) subst in - let exp = GTypRef (pinfo.pdata_type, params) in - let e = intern_rec_with_constraint env e exp in - let substf i = GTypVar subst.(i) in - let ret = subst_type substf pinfo.pdata_ptyp in - let r = intern_rec_with_constraint env r ret in - (GTacSet (pinfo.pdata_type, e, pinfo.pdata_indx, r), GTypRef (t_unit, [])) -| CTacExt (loc, ext) -> - let open Genintern in - let GenArg (Rawwit tag, _) = ext in - let tpe = interp_ml_object tag in - (** External objects do not have access to the named context because this is - not stable by dynamic semantics. *) - let genv = Global.env_of_context Environ.empty_named_context_val in - let ist = empty_glob_sign genv in - let ist = { ist with extra = Store.set ist.extra ltac2_env env } in - let (_, ext) = Flags.with_option Ltac_plugin.Tacintern.strict_check (fun () -> generic_intern ist ext) () in - (GTacExt ext, GTypRef (tpe.ml_type, [])) - -and intern_rec_with_constraint env e exp = - let loc = loc_of_tacexpr e in - let (e, t) = intern_rec env e in - let () = unify loc env t exp in - e - -and intern_let_rec env loc el e = - let fold accu ((loc, na), _, _) = match na with - | Anonymous -> accu - | Name id -> - if Id.Set.mem id accu then - user_err ~loc (str "Variable " ++ Id.print id ++ str " is bound several \ - times in this matching") - else Id.Set.add id accu - in - let _ = List.fold_left fold Id.Set.empty el in - let map env ((loc, na), t, e) = - let id = fresh_id env in - let env = push_name na (monomorphic (GTypVar id)) env in - (env, (loc, na, t, e, id)) - in - let (env, el) = List.fold_map map env el in - let fold (loc, na, tc, e, id) (el, tl) = - let loc_e = loc_of_tacexpr e in - let (e, t) = intern_rec env e in - let () = - if not (is_rec_rhs e) then - user_err ~loc:loc_e (str "This kind of expression is not allowed as \ - right-hand side of a recursive binding") - in - let () = unify loc env t (GTypVar id) in - let () = match tc with - | None -> () - | Some tc -> - let tc = intern_type env tc in - unify loc env t tc - in - ((na, e) :: el, t :: tl) - in - let (el, tl) = List.fold_right fold el ([], []) in - let (e, t) = intern_rec env e in - (GTacLet (true, el, e), t) - -(** For now, patterns recognized by the pattern-matching compiling are limited - to depth-one where leaves are either variables or catch-all *) -and intern_case env loc e pl = - let (e', t) = intern_rec env e in - let todo ~loc () = user_err ~loc (str "Pattern not handled yet") in - match get_pattern_kind env pl with - | PKind_any -> - let (pat, b) = List.hd pl in - let na = match intern_patexpr env pat with - | GPatVar na -> na - | _ -> assert false - in - let () = check_redundant_clause (List.tl pl) in - let env = push_name na (monomorphic t) env in - let (b, tb) = intern_rec env b in - (GTacLet (false, [na, e'], b), tb) - | PKind_empty -> - let kn = check_elt_empty loc env t in - let r = fresh_id env in - (GTacCse (e', GCaseAlg kn, [||], [||]), GTypVar r) - | PKind_tuple len -> - begin match pl with - | [] -> assert false - | [CPatTup (_, []), b] -> - let () = unify (loc_of_tacexpr e) env t (GTypRef (t_unit, [])) in - let (b, tb) = intern_rec env b in - (GTacCse (e', GCaseAlg t_unit, [|b|], [||]), tb) - | [CPatTup (_, pl), b] -> - let map = function - | CPatAny _ -> Anonymous - | CPatRef (loc, qid, []) -> - begin match get_constructor env qid with - | ArgVar (_, id) -> Name id - | ArgArg _ -> todo ~loc () - end - | p -> todo ~loc:(loc_of_patexpr p) () - in - let ids = Array.map_of_list map pl in - let tc = GTypTuple (List.map (fun _ -> GTypVar (fresh_id env)) pl) in - let () = unify (loc_of_tacexpr e) env t tc in - let (b, tb) = intern_rec env b in - (GTacCse (e', GCaseTuple len, [||], [|ids, b|]), tb) - | (p, _) :: _ -> todo ~loc:(loc_of_patexpr p) () - end - | PKind_variant kn -> - let subst, tc = fresh_reftype env kn in - let () = unify (loc_of_tacexpr e) env t tc in - let (params, def) = Tac2env.interp_type kn in - let cstrs = match def with - | GTydAlg c -> c - | _ -> assert false - in - let count (const, nonconst) (c, args) = match args with - | [] -> (succ const, nonconst) - | _ :: _ -> (const, succ nonconst) - in - let nconst, nnonconst = List.fold_left count (0, 0) cstrs in - let const = Array.make nconst None in - let nonconst = Array.make nnonconst None in - let ret = GTypVar (fresh_id env) in - let rec intern_branch = function - | [] -> () - | (pat, br) :: rem -> - let tbr = match pat with - | CPatAny _ -> - let () = check_redundant_clause rem in - let (br', brT) = intern_rec env br in - (** Fill all remaining branches *) - let fill (ncst, narg) (_, args) = - if List.is_empty args then - let () = - if Option.is_empty const.(ncst) then const.(ncst) <- Some br' - in - (succ ncst, narg) - else - let () = - if Option.is_empty const.(narg) then - let ids = Array.map_of_list (fun _ -> Anonymous) args in - nonconst.(narg) <- Some (ids, br') - in - (ncst, succ narg) - in - let _ = List.fold_left fill (0, 0) cstrs in - brT - | CPatRef (loc, qid, args) -> - let data = match get_constructor env qid with - | ArgVar _ -> todo ~loc () - | ArgArg (data, _) -> - let () = - let kn' = data.cdata_type in - if not (KerName.equal kn kn') then - invalid_pattern ~loc kn (GCaseAlg kn') - in - data - in - let get_id = function - | CPatAny _ -> Anonymous - | CPatRef (loc, qid, []) -> - begin match get_constructor env qid with - | ArgVar (_, id) -> Name id - | ArgArg _ -> todo ~loc () - end - | p -> todo ~loc:(loc_of_patexpr p) () - in - let ids = List.map get_id args in - let nids = List.length ids in - let nargs = List.length data.cdata_args in - let () = - if not (Int.equal nids nargs) then error_nargs_mismatch loc nargs nids - in - let fold env id tpe = - (** Instantiate all arguments *) - let subst n = GTypVar subst.(n) in - let tpe = subst_type subst tpe in - push_name id (monomorphic tpe) env - in - let nenv = List.fold_left2 fold env ids data.cdata_args in - let (br', brT) = intern_rec nenv br in - let () = - let index = match data.cdata_indx with - | Some i -> i - | None -> assert false - in - if List.is_empty args then - if Option.is_empty const.(index) then const.(index) <- Some br' - else warn_redundant_clause ~loc () - else - let ids = Array.of_list ids in - if Option.is_empty nonconst.(index) then nonconst.(index) <- Some (ids, br') - else warn_redundant_clause ~loc () - in - brT - | CPatTup (loc, tup) -> - invalid_pattern ~loc kn (GCaseTuple (List.length tup)) - in - let () = unify (loc_of_tacexpr br) env ret tbr in - intern_branch rem - in - let () = intern_branch pl in - let map = function - | None -> user_err ~loc (str "TODO: Unhandled match case") (** FIXME *) - | Some x -> x - in - let const = Array.map map const in - let nonconst = Array.map map nonconst in - let ce = GTacCse (e', GCaseAlg kn, const, nonconst) in - (ce, ret) - | PKind_open kn -> - let subst, tc = fresh_reftype env kn in - let () = unify (loc_of_tacexpr e) env t tc in - let ret = GTypVar (fresh_id env) in - let rec intern_branch map = function - | [] -> - user_err ~loc (str "Missing default case") - | (pat, br) :: rem -> - match intern_patexpr env pat with - | GPatVar na -> - let () = check_redundant_clause rem in - let nenv = push_name na (monomorphic tc) env in - let br' = intern_rec_with_constraint nenv br ret in - let def = (na, br') in - (map, def) - | GPatRef (knc, args) -> - let get = function - | GPatVar na -> na - | GPatRef _ | GPatTup _ -> - user_err ~loc (str "TODO: Unhandled match case") (** FIXME *) - in - let loc = loc_of_patexpr pat in - let ids = List.map get args in - let data = Tac2env.interp_constructor knc in - let () = - if not (KerName.equal kn data.cdata_type) then - invalid_pattern ~loc kn (GCaseAlg data.cdata_type) - in - let nids = List.length ids in - let nargs = List.length data.cdata_args in - let () = - if not (Int.equal nids nargs) then error_nargs_mismatch loc nargs nids - in - let fold env id tpe = - (** Instantiate all arguments *) - let subst n = GTypVar subst.(n) in - let tpe = subst_type subst tpe in - push_name id (monomorphic tpe) env - in - let nenv = List.fold_left2 fold env ids data.cdata_args in - let br' = intern_rec_with_constraint nenv br ret in - let map = - if KNmap.mem knc map then - let () = warn_redundant_clause ~loc () in - map - else - KNmap.add knc (Anonymous, Array.of_list ids, br') map - in - intern_branch map rem - | GPatTup tup -> - invalid_pattern ~loc kn (GCaseTuple (List.length tup)) - in - let (map, def) = intern_branch KNmap.empty pl in - (GTacWth { opn_match = e'; opn_branch = map; opn_default = def }, ret) - -and intern_constructor env loc kn args = - let cstr = interp_constructor kn in - let nargs = List.length cstr.cdata_args in - if Int.equal nargs (List.length args) then - let subst = Array.init cstr.cdata_prms (fun _ -> fresh_id env) in - let substf i = GTypVar subst.(i) in - let types = List.map (fun t -> subst_type substf t) cstr.cdata_args in - let ans = GTypRef (cstr.cdata_type, List.init cstr.cdata_prms (fun i -> GTypVar subst.(i))) in - let map arg tpe = intern_rec_with_constraint env arg tpe in - let args = List.map2 map args types in - match cstr.cdata_indx with - | Some idx -> - (GTacCst (GCaseAlg cstr.cdata_type, idx, args), ans) - | None -> - (GTacOpn (kn, args), ans) - else - error_nargs_mismatch loc nargs (List.length args) - -and intern_record env loc fs = - let map (proj, e) = - let loc = match proj with - | RelId (loc, _) -> loc - | AbsKn _ -> Loc.ghost - in - let proj = get_projection proj in - (loc, proj, e) - in - let fs = List.map map fs in - let kn = match fs with - | [] -> user_err ~loc (str "Cannot infer the corresponding record type") - | (_, proj, _) :: _ -> proj.pdata_type - in - let params, typdef = match Tac2env.interp_type kn with - | n, GTydRec def -> n, def - | _ -> assert false - in - let subst = Array.init params (fun _ -> fresh_id env) in - (** Set the answer [args] imperatively *) - let args = Array.make (List.length typdef) None in - let iter (loc, pinfo, e) = - if KerName.equal kn pinfo.pdata_type then - let index = pinfo.pdata_indx in - match args.(index) with - | None -> - let exp = subst_type (fun i -> GTypVar subst.(i)) pinfo.pdata_ptyp in - let e = intern_rec_with_constraint env e exp in - args.(index) <- Some e - | Some _ -> - let (name, _, _) = List.nth typdef pinfo.pdata_indx in - user_err ~loc (str "Field " ++ Id.print name ++ str " is defined \ - several times") - else - user_err ~loc (str "Field " ++ (*KerName.print knp ++*) str " does not \ - pertain to record definition " ++ pr_typref pinfo.pdata_type) - in - let () = List.iter iter fs in - let () = match Array.findi (fun _ o -> Option.is_empty o) args with - | None -> () - | Some i -> - let (field, _, _) = List.nth typdef i in - user_err ~loc (str "Field " ++ Id.print field ++ str " is undefined") - in - let args = Array.map_to_list Option.get args in - let tparam = List.init params (fun i -> GTypVar subst.(i)) in - (GTacCst (GCaseAlg kn, 0, args), GTypRef (kn, tparam)) - -let normalize env (count, vars) (t : UF.elt glb_typexpr) = - let get_var id = - try UF.Map.find id !vars - with Not_found -> - let () = assert env.env_opn in - let n = GTypVar !count in - let () = incr count in - let () = vars := UF.Map.add id n !vars in - n - in - let rec subst id = match UF.find id env.env_cst with - | id, None -> get_var id - | _, Some t -> subst_type subst t - in - subst_type subst t - -let intern e = - let env = empty_env () in - let (e, t) = intern_rec env e in - let count = ref 0 in - let vars = ref UF.Map.empty in - let t = normalize env (count, vars) t in - (e, (!count, t)) - -let intern_typedef self (ids, t) : glb_quant_typedef = - let env = { (empty_env ()) with env_rec = self } in - (** Initialize type parameters *) - let map id = get_alias id env in - let ids = List.map map ids in - let count = ref (List.length ids) in - let vars = ref UF.Map.empty in - let iter n id = vars := UF.Map.add id (GTypVar n) !vars in - let () = List.iteri iter ids in - (** Do not accept unbound type variables *) - let env = { env with env_opn = false } in - let intern t = - let t = intern_type env t in - normalize env (count, vars) t - in - let count = !count in - match t with - | CTydDef None -> (count, GTydDef None) - | CTydDef (Some t) -> (count, GTydDef (Some (intern t))) - | CTydAlg constrs -> - let map (c, t) = (c, List.map intern t) in - let constrs = List.map map constrs in - (count, GTydAlg constrs) - | CTydRec fields -> - let map (c, mut, t) = (c, mut, intern t) in - let fields = List.map map fields in - (count, GTydRec fields) - | CTydOpn -> (count, GTydOpn) - -let intern_open_type t = - let env = empty_env () in - let t = intern_type env t in - let count = ref 0 in - let vars = ref UF.Map.empty in - let t = normalize env (count, vars) t in - (!count, t) - -(** Globalization *) - -let add_name accu = function -| Name id -> Id.Set.add id accu -| Anonymous -> accu - -let get_projection0 var = match var with -| RelId (loc, qid) -> - let kn = try Tac2env.locate_projection qid with Not_found -> - user_err ~loc (pr_qualid qid ++ str " is not a projection") - in - kn -| AbsKn kn -> kn - -let rec globalize ids e = match e with -| CTacAtm _ -> e -| CTacRef ref -> - let mem id = Id.Set.mem id ids in - begin match get_variable0 mem ref with - | ArgVar _ -> e - | ArgArg kn -> CTacRef (AbsKn kn) - end -| CTacFun (loc, bnd, e) -> - let fold accu ((_, na), _) = add_name accu na in - let ids = List.fold_left fold ids bnd in - let e = globalize ids e in - CTacFun (loc, bnd, e) -| CTacApp (loc, e, el) -> - let e = globalize ids e in - let el = List.map (fun e -> globalize ids e) el in - CTacApp (loc, e, el) -| CTacLet (loc, isrec, bnd, e) -> - let fold accu ((_, na), _, _) = add_name accu na in - let ext = List.fold_left fold Id.Set.empty bnd in - let eids = Id.Set.union ext ids in - let e = globalize eids e in - let map (qid, t, e) = - let ids = if isrec then eids else ids in - (qid, t, globalize ids e) - in - let bnd = List.map map bnd in - CTacLet (loc, isrec, bnd, e) -| CTacTup (loc, el) -> - let el = List.map (fun e -> globalize ids e) el in - CTacTup (loc, el) -| CTacArr (loc, el) -> - let el = List.map (fun e -> globalize ids e) el in - CTacArr (loc, el) -| CTacLst (loc, el) -> - let el = List.map (fun e -> globalize ids e) el in - CTacLst (loc, el) -| CTacCnv (loc, e, t) -> - let e = globalize ids e in - CTacCnv (loc, e, t) -| CTacSeq (loc, e1, e2) -> - let e1 = globalize ids e1 in - let e2 = globalize ids e2 in - CTacSeq (loc, e1, e2) -| CTacCse (loc, e, bl) -> - let e = globalize ids e in - let bl = List.map (fun b -> globalize_case ids b) bl in - CTacCse (loc, e, bl) -| CTacRec (loc, r) -> - let map (p, e) = - let p = get_projection0 p in - let e = globalize ids e in - (AbsKn p, e) - in - CTacRec (loc, List.map map r) -| CTacPrj (loc, e, p) -> - let e = globalize ids e in - let p = get_projection0 p in - CTacPrj (loc, e, AbsKn p) -| CTacSet (loc, e, p, e') -> - let e = globalize ids e in - let p = get_projection0 p in - let e' = globalize ids e' in - CTacSet (loc, e, AbsKn p, e') -| CTacExt (loc, arg) -> - let arg = pr_argument_type (genarg_tag arg) in - CErrors.user_err ~loc (str "Cannot globalize generic arguments of type" ++ spc () ++ arg) - -and globalize_case ids (p, e) = - (globalize_pattern ids p, globalize ids e) - -and globalize_pattern ids p = match p with -| CPatAny _ -> p -| CPatRef (loc, cst, pl) -> - let cst = match get_constructor () cst with - | ArgVar _ -> cst - | ArgArg (_, knc) -> AbsKn knc - in - let pl = List.map (fun p -> globalize_pattern ids p) pl in - CPatRef (loc, cst, pl) -| CPatTup (loc, pl) -> - let pl = List.map (fun p -> globalize_pattern ids p) pl in - CPatTup (loc, pl) - -(** Kernel substitution *) - -open Mod_subst - -let rec subst_type subst t = match t with -| GTypVar _ -> t -| GTypArrow (t1, t2) -> - let t1' = subst_type subst t1 in - let t2' = subst_type subst t2 in - if t1' == t1 && t2' == t2 then t - else GTypArrow (t1', t2') -| GTypTuple tl -> - let tl'= List.smartmap (fun t -> subst_type subst t) tl in - if tl' == tl then t else GTypTuple tl' -| GTypRef (kn, tl) -> - let kn' = subst_kn subst kn in - let tl' = List.smartmap (fun t -> subst_type subst t) tl in - if kn' == kn && tl' == tl then t else GTypRef (kn', tl') - -let subst_case_info subst ci = match ci with -| GCaseAlg kn -> - let kn' = subst_kn subst kn in - if kn' == kn then ci else GCaseAlg kn' -| GCaseTuple _ -> ci - -let rec subst_expr subst e = match e with -| GTacAtm _ | GTacVar _ | GTacPrm _ -> e -| GTacRef kn -> GTacRef (subst_kn subst kn) -| GTacFun (ids, e) -> GTacFun (ids, subst_expr subst e) -| GTacApp (f, args) -> - GTacApp (subst_expr subst f, List.map (fun e -> subst_expr subst e) args) -| GTacLet (r, bs, e) -> - let bs = List.map (fun (na, e) -> (na, subst_expr subst e)) bs in - GTacLet (r, bs, subst_expr subst e) -| GTacArr el -> - GTacArr (List.map (fun e -> subst_expr subst e) el) -| GTacCst (t, n, el) as e0 -> - let t' = match t with - | GCaseAlg kn -> - let kn' = subst_kn subst kn in - if kn' == kn then t else GCaseAlg kn' - | GCaseTuple _ -> t - in - let el' = List.smartmap (fun e -> subst_expr subst e) el in - if t' == t && el' == el then e0 else GTacCst (t', n, el') -| GTacCse (e, ci, cse0, cse1) -> - let cse0' = Array.map (fun e -> subst_expr subst e) cse0 in - let cse1' = Array.map (fun (ids, e) -> (ids, subst_expr subst e)) cse1 in - let ci' = subst_case_info subst ci in - GTacCse (subst_expr subst e, ci', cse0', cse1') -| GTacWth { opn_match = e; opn_branch = br; opn_default = (na, def) } as e0 -> - let e' = subst_expr subst e in - let def' = subst_expr subst def in - let fold kn (self, vars, p) accu = - let kn' = subst_kn subst kn in - let p' = subst_expr subst p in - if kn' == kn && p' == p then accu - else KNmap.add kn' (self, vars, p') (KNmap.remove kn accu) - in - let br' = KNmap.fold fold br br in - if e' == e && br' == br && def' == def then e0 - else GTacWth { opn_match = e'; opn_default = (na, def'); opn_branch = br' } -| GTacPrj (kn, e, p) as e0 -> - let kn' = subst_kn subst kn in - let e' = subst_expr subst e in - if kn' == kn && e' == e then e0 else GTacPrj (kn', e', p) -| GTacSet (kn, e, p, r) as e0 -> - let kn' = subst_kn subst kn in - let e' = subst_expr subst e in - let r' = subst_expr subst r in - if kn' == kn && e' == e && r' == r then e0 else GTacSet (kn', e', p, r') -| GTacExt ext -> - let ext' = Genintern.generic_substitute subst ext in - if ext' == ext then e else GTacExt ext' -| GTacOpn (kn, el) as e0 -> - let kn' = subst_kn subst kn in - let el' = List.smartmap (fun e -> subst_expr subst e) el in - if kn' == kn && el' == el then e0 else GTacOpn (kn', el') - -let subst_typedef subst e = match e with -| GTydDef t -> - let t' = Option.smartmap (fun t -> subst_type subst t) t in - if t' == t then e else GTydDef t' -| GTydAlg constrs -> - let map (c, tl as p) = - let tl' = List.smartmap (fun t -> subst_type subst t) tl in - if tl' == tl then p else (c, tl') - in - let constrs' = List.smartmap map constrs in - if constrs' == constrs then e else GTydAlg constrs' -| GTydRec fields -> - let map (c, mut, t as p) = - let t' = subst_type subst t in - if t' == t then p else (c, mut, t') - in - let fields' = List.smartmap map fields in - if fields' == fields then e else GTydRec fields' -| GTydOpn -> GTydOpn - -let subst_quant_typedef subst (prm, def as qdef) = - let def' = subst_typedef subst def in - if def' == def then qdef else (prm, def') - -let subst_type_scheme subst (prm, t as sch) = - let t' = subst_type subst t in - if t' == t then sch else (prm, t') - -let subst_or_relid subst ref = match ref with -| RelId _ -> ref -| AbsKn kn -> - let kn' = subst_kn subst kn in - if kn' == kn then ref else AbsKn kn' - -let rec subst_rawtype subst t = match t with -| CTypVar _ -> t -| CTypArrow (loc, t1, t2) -> - let t1' = subst_rawtype subst t1 in - let t2' = subst_rawtype subst t2 in - if t1' == t1 && t2' == t2 then t else CTypArrow (loc, t1', t2') -| CTypTuple (loc, tl) -> - let tl' = List.smartmap (fun t -> subst_rawtype subst t) tl in - if tl' == tl then t else CTypTuple (loc, tl') -| CTypRef (loc, ref, tl) -> - let ref' = subst_or_relid subst ref in - let tl' = List.smartmap (fun t -> subst_rawtype subst t) tl in - if ref' == ref && tl' == tl then t else CTypRef (loc, ref', tl') - -let subst_tacref subst ref = match ref with -| RelId _ -> ref -| AbsKn (TacConstant kn) -> - let kn' = subst_kn subst kn in - if kn' == kn then ref else AbsKn (TacConstant kn') -| AbsKn (TacConstructor kn) -> - let kn' = subst_kn subst kn in - if kn' == kn then ref else AbsKn (TacConstructor kn') - -let subst_projection subst prj = match prj with -| RelId _ -> prj -| AbsKn kn -> - let kn' = subst_kn subst kn in - if kn' == kn then prj else AbsKn kn' - -let rec subst_rawpattern subst p = match p with -| CPatAny _ -> p -| CPatRef (loc, c, pl) -> - let pl' = List.smartmap (fun p -> subst_rawpattern subst p) pl in - let c' = match c with - | RelId _ -> c - | AbsKn kn -> - let kn' = subst_kn subst kn in - if kn' == kn then c else AbsKn kn' - in - if pl' == pl && c' == c then p else CPatRef (loc, c', pl') -| CPatTup (loc, pl) -> - let pl' = List.smartmap (fun p -> subst_rawpattern subst p) pl in - if pl' == pl then p else CPatTup (loc, pl') - -(** Used for notations *) -let rec subst_rawexpr subst t = match t with -| CTacAtm _ -> t -| CTacRef ref -> - let ref' = subst_tacref subst ref in - if ref' == ref then t else CTacRef ref' -| CTacFun (loc, bnd, e) -> - let map (na, t as p) = - let t' = Option.smartmap (fun t -> subst_rawtype subst t) t in - if t' == t then p else (na, t') - in - let bnd' = List.smartmap map bnd in - let e' = subst_rawexpr subst e in - if bnd' == bnd && e' == e then t else CTacFun (loc, bnd', e') -| CTacApp (loc, e, el) -> - let e' = subst_rawexpr subst e in - let el' = List.smartmap (fun e -> subst_rawexpr subst e) el in - if e' == e && el' == el then t else CTacApp (loc, e', el') -| CTacLet (loc, isrec, bnd, e) -> - let map (na, t, e as p) = - let t' = Option.smartmap (fun t -> subst_rawtype subst t) t in - let e' = subst_rawexpr subst e in - if t' == t && e' == e then p else (na, t', e') - in - let bnd' = List.smartmap map bnd in - let e' = subst_rawexpr subst e in - if bnd' == bnd && e' == e then t else CTacLet (loc, isrec, bnd', e') -| CTacTup (loc, el) -> - let el' = List.smartmap (fun e -> subst_rawexpr subst e) el in - if el' == el then t else CTacTup (loc, el') -| CTacArr (loc, el) -> - let el' = List.smartmap (fun e -> subst_rawexpr subst e) el in - if el' == el then t else CTacArr (loc, el') -| CTacLst (loc, el) -> - let el' = List.smartmap (fun e -> subst_rawexpr subst e) el in - if el' == el then t else CTacLst (loc, el') -| CTacCnv (loc, e, c) -> - let e' = subst_rawexpr subst e in - let c' = subst_rawtype subst c in - if c' == c && e' == e then t else CTacCnv (loc, e', c') -| CTacSeq (loc, e1, e2) -> - let e1' = subst_rawexpr subst e1 in - let e2' = subst_rawexpr subst e2 in - if e1' == e1 && e2' == e2 then t else CTacSeq (loc, e1', e2') -| CTacCse (loc, e, bl) -> - let map (p, e as x) = - let p' = subst_rawpattern subst p in - let e' = subst_rawexpr subst e in - if p' == p && e' == e then x else (p', e') - in - let e' = subst_rawexpr subst e in - let bl' = List.smartmap map bl in - if e' == e && bl' == bl then t else CTacCse (loc, e', bl') -| CTacRec (loc, el) -> - let map (prj, e as p) = - let prj' = subst_projection subst prj in - let e' = subst_rawexpr subst e in - if prj' == prj && e' == e then p else (prj', e') - in - let el' = List.smartmap map el in - if el' == el then t else CTacRec (loc, el') -| CTacPrj (loc, e, prj) -> - let prj' = subst_projection subst prj in - let e' = subst_rawexpr subst e in - if prj' == prj && e' == e then t else CTacPrj (loc, e', prj') -| CTacSet (loc, e, prj, r) -> - let prj' = subst_projection subst prj in - let e' = subst_rawexpr subst e in - let r' = subst_rawexpr subst r in - if prj' == prj && e' == e && r' == r then t else CTacSet (loc, e', prj', r') -| CTacExt _ -> assert false (** Should not be generated by gloabalization *) - -(** Registering *) - -let () = - let open Genintern in - let intern ist tac = - let env = match Genintern.Store.get ist.extra ltac2_env with - | None -> empty_env () - | Some env -> env - in - let loc = loc_of_tacexpr tac in - let (tac, t) = intern_rec env tac in - let () = check_elt_unit loc env t in - (ist, tac) - in - Genintern.register_intern0 wit_ltac2 intern -let () = Genintern.register_subst0 wit_ltac2 subst_expr diff --git a/tac2intern.mli b/tac2intern.mli deleted file mode 100644 index 3d400a5cdd..0000000000 --- a/tac2intern.mli +++ /dev/null @@ -1,41 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* Loc.t - -val intern : raw_tacexpr -> glb_tacexpr * type_scheme -val intern_typedef : (KerName.t * int) Id.Map.t -> raw_quant_typedef -> glb_quant_typedef -val intern_open_type : raw_typexpr -> type_scheme - -(** Check that a term is a value. Only values are safe to marshall between - processes. *) -val is_value : glb_tacexpr -> bool -val check_unit : ?loc:Loc.t -> int glb_typexpr -> unit - -val subst_type : substitution -> 'a glb_typexpr -> 'a glb_typexpr -val subst_expr : substitution -> glb_tacexpr -> glb_tacexpr -val subst_quant_typedef : substitution -> glb_quant_typedef -> glb_quant_typedef -val subst_type_scheme : substitution -> type_scheme -> type_scheme - -val subst_rawexpr : substitution -> raw_tacexpr -> raw_tacexpr - -(** {5 Notations} *) - -val globalize : Id.Set.t -> raw_tacexpr -> raw_tacexpr -(** Replaces all qualified identifiers by their corresponding kernel name. The - set represents bound variables in the context. *) - -(** Errors *) - -val error_nargs_mismatch : Loc.t -> int -> int -> 'a -val error_nparams_mismatch : Loc.t -> int -> int -> 'a diff --git a/tac2interp.ml b/tac2interp.ml deleted file mode 100644 index 664b7de3d6..0000000000 --- a/tac2interp.ml +++ /dev/null @@ -1,160 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* - let c = Tac2print.pr_constructor kn in - hov 0 (str "Uncaught Ltac2 exception:" ++ spc () ++ hov 0 c) -| _ -> raise Unhandled -end - -let val_exn = Geninterp.Val.create "ltac2:exn" - -type environment = valexpr Id.Map.t - -let empty_environment = Id.Map.empty - -let push_name ist id v = match id with -| Anonymous -> ist -| Name id -> Id.Map.add id v ist - -let get_var ist id = - try Id.Map.find id ist with Not_found -> - anomaly (str "Unbound variable " ++ Id.print id) - -let get_ref ist kn = - try pi2 (Tac2env.interp_global kn) with Not_found -> - anomaly (str "Unbound reference" ++ KerName.print kn) - -let return = Proofview.tclUNIT - -let rec interp ist = function -| GTacAtm (AtmInt n) -> return (ValInt n) -| GTacAtm (AtmStr s) -> return (ValStr (Bytes.of_string s)) -| GTacVar id -> return (get_var ist id) -| GTacRef qid -> return (get_ref ist qid) -| GTacFun (ids, e) -> - let cls = { clos_env = ist; clos_var = ids; clos_exp = e } in - return (ValCls cls) -| GTacApp (f, args) -> - interp ist f >>= fun f -> - Proofview.Monad.List.map (fun e -> interp ist e) args >>= fun args -> - interp_app f args -| GTacLet (false, el, e) -> - let fold accu (na, e) = - interp ist e >>= fun e -> - return (push_name accu na e) - in - Proofview.Monad.List.fold_left fold ist el >>= fun ist -> - interp ist e -| GTacLet (true, el, e) -> - let map (na, e) = match e with - | GTacFun (ids, e) -> - let cls = { clos_env = ist; clos_var = ids; clos_exp = e } in - na, cls - | _ -> anomaly (str "Ill-formed recursive function") - in - let fixs = List.map map el in - let fold accu (na, cls) = match na with - | Anonymous -> accu - | Name id -> Id.Map.add id (ValCls cls) accu - in - let ist = List.fold_left fold ist fixs in - (** Hack to make a cycle imperatively in the environment *) - let iter (_, e) = e.clos_env <- ist in - let () = List.iter iter fixs in - interp ist e -| GTacArr el -> - Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> - return (ValBlk (0, Array.of_list el)) -| GTacCst (_, n, []) -> return (ValInt n) -| GTacCst (_, n, el) -> - Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> - return (ValBlk (n, Array.of_list el)) -| GTacCse (e, _, cse0, cse1) -> - interp ist e >>= fun e -> interp_case ist e cse0 cse1 -| GTacWth { opn_match = e; opn_branch = cse; opn_default = def } -> - interp ist e >>= fun e -> interp_with ist e cse def -| GTacPrj (_, e, p) -> - interp ist e >>= fun e -> interp_proj ist e p -| GTacSet (_, e, p, r) -> - interp ist e >>= fun e -> - interp ist r >>= fun r -> - interp_set ist e p r -| GTacOpn (kn, el) -> - Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> - return (ValOpn (kn, Array.of_list el)) -| GTacPrm (ml, el) -> - Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> - Tac2env.interp_primitive ml el -| GTacExt e -> - let GenArg (Glbwit tag, e) = e in - let tpe = Tac2env.interp_ml_object tag in - tpe.Tac2env.ml_interp ist e >>= fun e -> return (ValExt e) - -and interp_app f args = match f with -| ValCls { clos_env = ist; clos_var = ids; clos_exp = e } -> - let rec push ist ids args = match ids, args with - | [], [] -> interp ist e - | [], _ :: _ -> interp ist e >>= fun f -> interp_app f args - | _ :: _, [] -> - let cls = { clos_env = ist; clos_var = ids; clos_exp = e } in - return (ValCls cls) - | id :: ids, arg :: args -> push (push_name ist id arg) ids args - in - push ist ids args -| ValExt _ | ValInt _ | ValBlk _ | ValStr _ | ValOpn _ -> - anomaly (str "Unexpected value shape") - -and interp_case ist e cse0 cse1 = match e with -| ValInt n -> interp ist cse0.(n) -| ValBlk (n, args) -> - let (ids, e) = cse1.(n) in - let ist = CArray.fold_left2 push_name ist ids args in - interp ist e -| ValExt _ | ValStr _ | ValCls _ | ValOpn _ -> - anomaly (str "Unexpected value shape") - -and interp_with ist e cse def = match e with -| ValOpn (kn, args) -> - let br = try Some (KNmap.find kn cse) with Not_found -> None in - begin match br with - | None -> - let (self, def) = def in - let ist = push_name ist self e in - interp ist def - | Some (self, ids, p) -> - let ist = push_name ist self e in - let ist = CArray.fold_left2 push_name ist ids args in - interp ist p - end -| ValInt _ | ValBlk _ | ValExt _ | ValStr _ | ValCls _ -> - anomaly (str "Unexpected value shape") - -and interp_proj ist e p = match e with -| ValBlk (_, args) -> - return args.(p) -| ValInt _ | ValExt _ | ValStr _ | ValCls _ | ValOpn _ -> - anomaly (str "Unexpected value shape") - -and interp_set ist e p r = match e with -| ValBlk (_, args) -> - let () = args.(p) <- r in - return (ValInt 0) -| ValInt _ | ValExt _ | ValStr _ | ValCls _ | ValOpn _ -> - anomaly (str "Unexpected value shape") diff --git a/tac2interp.mli b/tac2interp.mli deleted file mode 100644 index bf6b2d4dde..0000000000 --- a/tac2interp.mli +++ /dev/null @@ -1,28 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* glb_tacexpr -> valexpr Proofview.tactic - -val interp_app : valexpr -> valexpr list -> valexpr Proofview.tactic - -(** {5 Exceptions} *) - -exception LtacError of KerName.t * valexpr array -(** Ltac2-defined exceptions seen from OCaml side *) - -val val_exn : Exninfo.iexn Geninterp.Val.typ -(** Toplevel representation of OCaml exceptions. Invariant: no [LtacError] - should be put into a value with tag [val_exn]. *) diff --git a/tac2print.ml b/tac2print.ml deleted file mode 100644 index e6f0582e3d..0000000000 --- a/tac2print.ml +++ /dev/null @@ -1,296 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* str "'" ++ str (pr n) - | GTypRef (kn, []) -> pr_typref kn - | GTypRef (kn, [t]) -> - let paren = match lvl with - | T5_r | T5_l | T2 | T1 -> fun x -> x - | T0 -> paren - in - paren (pr_glbtype lvl t ++ spc () ++ pr_typref kn) - | GTypRef (kn, tl) -> - let paren = match lvl with - | T5_r | T5_l | T2 | T1 -> fun x -> x - | T0 -> paren - in - paren (str "(" ++ prlist_with_sep (fun () -> str ", ") (pr_glbtype lvl) tl ++ str ")" ++ spc () ++ pr_typref kn) - | GTypArrow (t1, t2) -> - let paren = match lvl with - | T5_r -> fun x -> x - | T5_l | T2 | T1 | T0 -> paren - in - paren (pr_glbtype T5_l t1 ++ spc () ++ str "->" ++ spc () ++ pr_glbtype T5_r t2) - | GTypTuple tl -> - let paren = match lvl with - | T5_r | T5_l -> fun x -> x - | T2 | T1 | T0 -> paren - in - paren (prlist_with_sep (fun () -> str " * ") (pr_glbtype T2) tl) - in - hov 0 (pr_glbtype lvl c) - -let pr_glbtype pr c = pr_glbtype_gen pr T5_r c - -let int_name () = - let vars = ref Int.Map.empty in - fun n -> - if Int.Map.mem n !vars then Int.Map.find n !vars - else - let num = Int.Map.cardinal !vars in - let base = num mod 26 in - let rem = num / 26 in - let name = String.make 1 (Char.chr (97 + base)) in - let suff = if Int.equal rem 0 then "" else string_of_int rem in - let name = name ^ suff in - let () = vars := Int.Map.add n name !vars in - name - -(** Term printing *) - -let pr_constructor kn = - Libnames.pr_qualid (Tac2env.shortest_qualid_of_ltac (TacConstructor kn)) - -let pr_projection kn = - Libnames.pr_qualid (Tac2env.shortest_qualid_of_projection kn) - -type exp_level = Tac2expr.exp_level = -| E5 -| E4 -| E3 -| E2 -| E1 -| E0 - -let pr_atom = function -| AtmInt n -> int n -| AtmStr s -> qstring s - -let pr_name = function -| Name id -> Id.print id -| Anonymous -> str "_" - -let find_constructor n empty def = - let rec find n = function - | [] -> assert false - | (id, []) :: rem -> - if empty then - if Int.equal n 0 then id - else find (pred n) rem - else find n rem - | (id, _ :: _) :: rem -> - if not empty then - if Int.equal n 0 then id - else find (pred n) rem - else find n rem - in - find n def - -let order_branches cbr nbr def = - let rec order cidx nidx def = match def with - | [] -> [] - | (id, []) :: rem -> - let ans = order (succ cidx) nidx rem in - (id, [], cbr.(cidx)) :: ans - | (id, _ :: _) :: rem -> - let ans = order cidx (succ nidx) rem in - let (vars, e) = nbr.(nidx) in - (id, Array.to_list vars, e) :: ans - in - order 0 0 def - -let pr_glbexpr_gen lvl c = - let rec pr_glbexpr lvl = function - | GTacAtm atm -> pr_atom atm - | GTacVar id -> Id.print id - | GTacRef gr -> - let qid = shortest_qualid_of_ltac (TacConstant gr) in - Libnames.pr_qualid qid - | GTacFun (nas, c) -> - let nas = pr_sequence pr_name nas in - let paren = match lvl with - | E0 | E1 | E2 | E3 | E4 -> paren - | E5 -> fun x -> x - in - paren (str "fun" ++ spc () ++ nas ++ spc () ++ str "=>" ++ spc () ++ - hov 0 (pr_glbexpr E5 c)) - | GTacApp (c, cl) -> - let paren = match lvl with - | E0 -> paren - | E1 | E2 | E3 | E4 | E5 -> fun x -> x - in - paren (pr_glbexpr E1 c ++ spc () ++ (pr_sequence (pr_glbexpr E0) cl)) - | GTacLet (mut, bnd, e) -> - let paren = match lvl with - | E0 | E1 | E2 | E3 | E4 -> paren - | E5 -> fun x -> x - in - let mut = if mut then str "rec" ++ spc () else mt () in - let pr_bnd (na, e) = - pr_name na ++ spc () ++ str ":=" ++ spc () ++ hov 2 (pr_glbexpr E5 e) ++ spc () - in - let bnd = prlist_with_sep (fun () -> str "with" ++ spc ()) pr_bnd bnd in - paren (str "let" ++ spc () ++ mut ++ bnd ++ str "in" ++ spc () ++ pr_glbexpr E5 e) - | GTacCst (GCaseTuple _, _, cl) -> - let paren = match lvl with - | E0 | E1 -> paren - | E2 | E3 | E4 | E5 -> fun x -> x - in - paren (prlist_with_sep (fun () -> str "," ++ spc ()) (pr_glbexpr E1) cl) - | GTacArr cl -> - mt () (** FIXME when implemented *) - | GTacCst (GCaseAlg tpe, n, cl) -> - begin match Tac2env.interp_type tpe with - | _, GTydAlg def -> - let paren = match lvl with - | E0 -> paren - | E1 | E2 | E3 | E4 | E5 -> fun x -> x - in - let id = find_constructor n (List.is_empty cl) def in - let kn = change_kn_label tpe id in - let cl = match cl with - | [] -> mt () - | _ -> spc () ++ pr_sequence (pr_glbexpr E0) cl - in - paren (pr_constructor kn ++ cl) - | _, GTydRec def -> - let args = List.combine def cl in - let pr_arg ((id, _, _), arg) = - let kn = change_kn_label tpe id in - pr_projection kn ++ spc () ++ str ":=" ++ spc () ++ pr_glbexpr E1 arg - in - let args = prlist_with_sep (fun () -> str ";" ++ spc ()) pr_arg args in - str "{" ++ spc () ++ args ++ spc () ++ str "}" - | _, (GTydDef _ | GTydOpn) -> assert false - end - | GTacCse (e, info, cst_br, ncst_br) -> - let e = pr_glbexpr E5 e in - let br = match info with - | GCaseAlg kn -> - let def = match Tac2env.interp_type kn with - | _, GTydAlg def -> def - | _, GTydDef _ | _, GTydRec _ | _, GTydOpn -> assert false - in - let br = order_branches cst_br ncst_br def in - let pr_branch (cstr, vars, p) = - let cstr = change_kn_label kn cstr in - let cstr = pr_constructor cstr in - let vars = match vars with - | [] -> mt () - | _ -> spc () ++ pr_sequence pr_name vars - in - hov 0 (str "|" ++ spc () ++ cstr ++ vars ++ spc () ++ str "=>" ++ spc () ++ - hov 2 (pr_glbexpr E5 p)) ++ spc () - in - prlist pr_branch br - | GCaseTuple n -> - let (vars, p) = ncst_br.(0) in - let p = pr_glbexpr E5 p in - let vars = prvect_with_sep (fun () -> str "," ++ spc ()) pr_name vars in - str "|" ++ spc () ++ paren vars ++ spc () ++ str "=>" ++ spc () ++ p - in - hov 0 (hov 0 (str "match" ++ spc () ++ e ++ spc () ++ str "with") ++ spc () ++ Pp.v 0 br ++ str "end") - | GTacWth wth -> - let e = pr_glbexpr E5 wth.opn_match in - let pr_pattern c self vars p = - let self = match self with - | Anonymous -> mt () - | Name id -> spc () ++ str "as" ++ spc () ++ Id.print id - in - hov 0 (str "|" ++ spc () ++ c ++ vars ++ self ++ spc () ++ str "=>" ++ spc () ++ - hov 2 (pr_glbexpr E5 p)) ++ spc () - in - let pr_branch (cstr, (self, vars, p)) = - let cstr = pr_constructor cstr in - let vars = match Array.to_list vars with - | [] -> mt () - | vars -> spc () ++ pr_sequence pr_name vars - in - pr_pattern cstr self vars p - in - let br = prlist pr_branch (KNmap.bindings wth.opn_branch) in - let (def_as, def_p) = wth.opn_default in - let def = pr_pattern (str "_") def_as (mt ()) def_p in - let br = br ++ def in - hov 0 (hov 0 (str "match" ++ spc () ++ e ++ spc () ++ str "with") ++ spc () ++ Pp.v 0 br ++ str "end") - | GTacPrj (kn, e, n) -> - let def = match Tac2env.interp_type kn with - | _, GTydRec def -> def - | _, GTydDef _ | _, GTydAlg _ | _, GTydOpn -> assert false - in - let (proj, _, _) = List.nth def n in - let proj = change_kn_label kn proj in - let proj = pr_projection proj in - let e = pr_glbexpr E0 e in - e ++ str "." ++ paren proj - | GTacSet (kn, e, n, r) -> - let def = match Tac2env.interp_type kn with - | _, GTydRec def -> def - | _, GTydDef _ | _, GTydAlg _ | _, GTydOpn -> assert false - in - let (proj, _, _) = List.nth def n in - let proj = change_kn_label kn proj in - let proj = pr_projection proj in - let e = pr_glbexpr E0 e in - let r = pr_glbexpr E1 r in - e ++ str "." ++ paren proj ++ spc () ++ str ":=" ++ spc () ++ r - | GTacOpn (kn, cl) -> - let paren = match lvl with - | E0 -> paren - | E1 | E2 | E3 | E4 | E5 -> fun x -> x - in - let c = pr_constructor kn in - paren (c ++ spc () ++ (pr_sequence (pr_glbexpr E0) cl)) - | GTacExt arg -> - let GenArg (Glbwit tag, arg) = arg in - let name = match tag with - | ExtraArg tag -> ArgT.repr tag - | _ -> assert false - in - str name ++ str ":" ++ paren (Genprint.glb_print tag arg) - | GTacPrm (prm, args) -> - let args = match args with - | [] -> mt () - | _ -> spc () ++ pr_sequence (pr_glbexpr E0) args - in - str "@external" ++ spc () ++ qstring prm.mltac_plugin ++ spc () ++ - qstring prm.mltac_tactic ++ args - in - hov 0 (pr_glbexpr lvl c) - -let pr_glbexpr c = - pr_glbexpr_gen E5 c diff --git a/tac2print.mli b/tac2print.mli deleted file mode 100644 index ddd599641d..0000000000 --- a/tac2print.mli +++ /dev/null @@ -1,37 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* std_ppcmds -val pr_glbtype_gen : ('a -> string) -> typ_level -> 'a glb_typexpr -> std_ppcmds -val pr_glbtype : ('a -> string) -> 'a glb_typexpr -> std_ppcmds - -(** {5 Printing expressions} *) - -val pr_constructor : ltac_constructor -> std_ppcmds -val pr_projection : ltac_projection -> std_ppcmds -val pr_glbexpr_gen : exp_level -> glb_tacexpr -> std_ppcmds -val pr_glbexpr : glb_tacexpr -> std_ppcmds - -(** {5 Utilities} *) - -val int_name : unit -> (int -> string) -(** Create a function that give names to integers. The names are generated on - the fly, in the order they are encountered. *) diff --git a/theories/Array.v b/theories/Array.v new file mode 100644 index 0000000000..11b64e3515 --- /dev/null +++ b/theories/Array.v @@ -0,0 +1,14 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* 'a -> 'a array := "ltac2" "array_make". +Ltac2 @external length : 'a array -> int := "ltac2" "array_length". +Ltac2 @external get : 'a array -> int -> 'a := "ltac2" "array_get". +Ltac2 @external set : 'a array -> int -> 'a -> unit := "ltac2" "array_set". diff --git a/theories/Constr.v b/theories/Constr.v new file mode 100644 index 0000000000..c340e3aa87 --- /dev/null +++ b/theories/Constr.v @@ -0,0 +1,43 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* constr := "ltac2" "constr_type". +(** Return the type of a term *) + +Ltac2 @ external equal : constr -> constr -> bool := "ltac2" "constr_equal". +(** Strict syntactic equality: only up to α-conversion and evar expansion *) + +Module Unsafe. + +(** Low-level access to kernel term. Use with care! *) + +Ltac2 Type kind := [ +| Rel (int) +| Var (ident) +| Meta (meta) +| Evar (evar, constr list) +| Sort (sort) +| Cast (constr, cast, constr) +| Prod (ident option, constr, constr) +| Lambda (ident option, constr, constr) +| LetIn (ident option, constr, constr, constr) +| App (constr, constr list) +| Constant (constant, instance) +| Ind (inductive, instance) +| Constructor (inductive, instance) +(* + | Case of case_info * 'constr * 'constr * 'constr array + | Fix of ('constr, 'types) pfixpoint + | CoFix of ('constr, 'types) pcofixpoint +*) +| Proj (projection, constr) +]. + +End Unsafe. diff --git a/theories/Control.v b/theories/Control.v new file mode 100644 index 0000000000..3bc572547c --- /dev/null +++ b/theories/Control.v @@ -0,0 +1,49 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* 'a := "ltac2" "throw". +(** Fatal exception throwing. This does not induce backtracking. *) + +(** Generic backtracking control *) + +Ltac2 @ external zero : exn -> 'a := "ltac2" "zero". +Ltac2 @ external plus : (unit -> 'a) -> (exn -> 'a) -> 'a := "ltac2" "plus". +Ltac2 @ external once : (unit -> 'a) -> 'a := "ltac2" "once". +Ltac2 @ external dispatch : (unit -> unit) list -> unit := "ltac2" "dispatch". +Ltac2 @ external extend : (unit -> unit) list -> (unit -> unit) -> (unit -> unit) list -> unit := "ltac2" "extend". +Ltac2 @ external enter : (unit -> unit) -> unit := "ltac2" "enter". + +(** Proof state manipulation *) + +Ltac2 @ external focus : int -> int -> (unit -> 'a) -> 'a := "ltac2" "focus". +Ltac2 @ external shelve : unit -> unit := "ltac2" "shelve". +Ltac2 @ external shelve_unifiable : unit -> unit := "ltac2" "shelve_unifiable". + +Ltac2 @ external new_goal : evar -> unit := "ltac2" "new_goal". +(** Adds the given evar to the list of goals as the last one. If it is + already defined in the current state, don't do anything. Panics if the + evar is not in the current state. *) + +(** Goal inspection *) + +Ltac2 @ external goal : unit -> constr := "ltac2" "goal". +(** Panics if there is not exactly one goal under focus. Otherwise returns + the conclusion of this goal. *) + +Ltac2 @ external hyp : ident -> constr := "ltac2" "hyp". +(** Panics if there is more than one goal under focus. If there is no + goal under focus, looks for the section variable with the given name. + If there is one, looks for the hypothesis with the given name. *) + +(** Refinement *) + +Ltac2 @ external refine : (unit -> constr) -> unit := "ltac2" "refine". diff --git a/theories/Init.v b/theories/Init.v new file mode 100644 index 0000000000..1d2d40f5c0 --- /dev/null +++ b/theories/Init.v @@ -0,0 +1,56 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* int -> bool := "ltac2" "int_equal". +Ltac2 @ external compare : int -> int -> int := "ltac2" "int_compare". +Ltac2 @ external add : int -> int -> int := "ltac2" "int_add". +Ltac2 @ external sub : int -> int -> int := "ltac2" "int_sub". +Ltac2 @ external mul : int -> int -> int := "ltac2" "int_mul". +Ltac2 @ external neg : int -> int := "ltac2" "int_neg". diff --git a/theories/Ltac2.v b/theories/Ltac2.v new file mode 100644 index 0000000000..221f7be424 --- /dev/null +++ b/theories/Ltac2.v @@ -0,0 +1,16 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* unit := "ltac2" "print". + +Ltac2 @ external of_string : string -> message := "ltac2" "message_of_string". + +Ltac2 @ external of_int : int -> message := "ltac2" "message_of_int". + +Ltac2 @ external of_constr : constr -> message := "ltac2" "message_of_constr". +(** Panics if there is more than one goal under focus. *) + +Ltac2 @ external concat : message -> message -> message := "ltac2" "message_concat". diff --git a/theories/String.v b/theories/String.v new file mode 100644 index 0000000000..99e1dab76b --- /dev/null +++ b/theories/String.v @@ -0,0 +1,14 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* char -> string := "ltac2" "string_make". +Ltac2 @external length : string -> int := "ltac2" "string_length". +Ltac2 @external get : string -> int -> char := "ltac2" "string_get". +Ltac2 @external set : string -> int -> char -> unit := "ltac2" "string_set". -- cgit v1.2.3 From fed8f69b682c03ff901966890efd2d9d3ea91004 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 24 Jul 2017 18:40:27 +0200 Subject: Fix library hardwired prefix. --- src/tac2env.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/tac2env.ml b/src/tac2env.ml index 5ccdd018ee..08c7b321be 100644 --- a/src/tac2env.ml +++ b/src/tac2env.ml @@ -235,7 +235,7 @@ let interp_ml_object t = MLType.obj t (** Absolute paths *) let coq_prefix = - MPfile (DirPath.make (List.map Id.of_string ["Init"; "ltac2"; "Coq"])) + MPfile (DirPath.make (List.map Id.of_string ["Init"; "Ltac2"])) (** Generic arguments *) -- cgit v1.2.3 From 484cf6add4aeb5faaa90f716d5acd2cc2bdf13b3 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 24 Jul 2017 18:38:32 +0200 Subject: Adding quick-n-dirty tests. --- tests/ltac2.v | 119 ++++++++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 119 insertions(+) create mode 100644 tests/ltac2.v diff --git a/tests/ltac2.v b/tests/ltac2.v new file mode 100644 index 0000000000..3d3d0032c5 --- /dev/null +++ b/tests/ltac2.v @@ -0,0 +1,119 @@ +Require Import Ltac2.Ltac2. + +Ltac2 foo (_ : int) := + let f (x : int) := x in + let _ := f 0 in + f 1. + +Print Ltac2 foo. + +Import Control. + +Ltac2 exact x := refine (fun _ => x). + +Print Ltac2 refine. +Print Ltac2 exact. + +Ltac2 foo' _ := ident:(bla). + +Print Ltac2 foo'. + +Ltac2 bar x H := match x with +| None => constr:(fun H => ltac2:(exact (hyp ident:(H))) -> nat) +| Some x => x +end. + +Print Ltac2 bar. + +Ltac2 qux := Some 0. + +Print Ltac2 qux. + +Ltac2 Type foo := [ Foo (int) ]. + +Fail Ltac2 qux0 := Foo None. + +Ltac2 Type 'a ref := { mutable contents : 'a }. + +Fail Ltac2 qux0 := { contents := None }. +Ltac2 foo0 _ := { contents := None }. + +Print Ltac2 foo0. + +Ltac2 qux0 x := x.(contents). +Ltac2 qux1 x := x.(contents) := x.(contents). + +Ltac2 qux2 := ([1;2], true). + +Print Ltac2 qux0. +Print Ltac2 qux1. +Print Ltac2 qux2. + +Import Control. + +Ltac2 qux3 x := constr:(nat -> ltac2:(refine (fun _ => hyp x))). + +Print Ltac2 qux3. + +Ltac2 qux4 f x := x, (f x, x). + +Print Ltac2 qux4. + +Ltac2 Type rec nat := [ O | S (nat) ]. + +Ltac2 message_of_nat n := +let rec aux n := +match n with +| O => Message.of_string "O" +| S n => Message.concat (Message.of_string "S") (aux n) +end in aux n. + +Print Ltac2 message_of_nat. + +Ltac2 numgoals (_ : unit) := + let r := { contents := O } in + enter (fun _ => r.(contents) := S (r.(contents))); + r.(contents). + +Print Ltac2 numgoals. + +Goal True /\ False. +Proof. +let n := numgoals () in Message.print (message_of_nat n). +refine (fun _ => open_constr:((fun x => conj _ _) 0)); (). +let n := numgoals () in Message.print (message_of_nat n). + +Fail (hyp ident:(x)). +Fail (enter (fun _ => hyp ident:(There_is_no_spoon); ())). + +enter (fun _ => Message.print (Message.of_string "foo")). + +enter (fun _ => Message.print (Message.of_constr (goal ()))). +Fail enter (fun _ => Message.print (Message.of_constr (qux3 ident:(x)))). +enter (fun _ => plus (fun _ => constr:(_); ()) (fun _ => ())). +plus + (fun _ => enter (fun _ => let x := ident:(foo) in let _ := hyp x in ())) (fun _ => Message.print (Message.of_string "failed")). +let x := { contents := 0 } in +let x := x.(contents) := x.(contents) in x. +Abort. + +Ltac2 Type exn ::= [ Foo ]. + +Goal True. +Proof. +plus (fun _ => zero Foo) (fun _ => ()). +Abort. + +Ltac2 Type exn ::= [ Bar (string) ]. + +Goal True. +Proof. +Fail zero (Bar "lol"). +Abort. + +Ltac2 Notation "refine!" c(constr) := refine c. + +Goal True. +Proof. +refine! I. +Abort. -- cgit v1.2.3 From 2c9233d40d75492fa7f06c09c2bf4f7b16fd1280 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 24 Jul 2017 18:49:10 +0200 Subject: Filling the README. --- README.md | 23 +++++++++++++++++++++++ 1 file changed, 23 insertions(+) diff --git a/README.md b/README.md index e69de29bb2..a5aa40eeb6 100644 --- a/README.md +++ b/README.md @@ -0,0 +1,23 @@ +Overview +======== + +This is a standalone version of the Ltac2 plugin. Ltac2 is an attempt at +providing the Coq users with a tactic language that is more robust and more +expressive than the venerable Ltac langue. + +Status +======== + +It is mostly a toy to experiment for now, and the implementation is quite +bug-ridden. Don't mistake this for a final product! + +Installation +============ + +This should compile with Coq 8.7, assuming the `COQLIB` variable is set +correctly. Standard procedures for `coq_makefile`-generated plugins apply. + +Demo +==== + +Horrible test-files are provided in the `tests` folder. Not for kids. -- cgit v1.2.3 From 5db849e1aa0d543d31389f5b10b6d863fcabce09 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 24 Jul 2017 18:49:19 +0200 Subject: Removing a spurious file. --- vo.itarget | 8 -------- 1 file changed, 8 deletions(-) delete mode 100644 vo.itarget diff --git a/vo.itarget b/vo.itarget deleted file mode 100644 index 5777585681..0000000000 --- a/vo.itarget +++ /dev/null @@ -1,8 +0,0 @@ -Init.vo -Int.vo -String.vo -Array.vo -Control.vo -Message.vo -Constr.vo -Ltac2.vo -- cgit v1.2.3 From 1fe83d5d791f0ff91ffa032b556f02c6cc4cbfed Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 24 Jul 2017 18:54:02 +0200 Subject: Fix typo. --- README.md | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/README.md b/README.md index a5aa40eeb6..4d6879d8af 100644 --- a/README.md +++ b/README.md @@ -3,7 +3,7 @@ Overview This is a standalone version of the Ltac2 plugin. Ltac2 is an attempt at providing the Coq users with a tactic language that is more robust and more -expressive than the venerable Ltac langue. +expressive than the venerable Ltac language. Status ======== -- cgit v1.2.3 From a647c38d3024f34711fbaa66975b5812097c33cc Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 24 Jul 2017 19:29:09 +0200 Subject: Properly handle parsing of list patterns. --- src/g_ltac2.ml4 | 6 +++++- 1 file changed, 5 insertions(+), 1 deletion(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 36057b3a67..6cdbccb11d 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -33,7 +33,11 @@ GEXTEND Gram tac2pat: [ "1" LEFTA [ id = Prim.qualid; pl = LIST1 tac2pat LEVEL "0" -> CPatRef (!@loc, RelId id, pl) - | id = Prim.qualid -> CPatRef (!@loc, RelId id, []) ] + | id = Prim.qualid -> CPatRef (!@loc, RelId id, []) + | "["; "]" -> CPatRef (!@loc, AbsKn Tac2core.Core.c_nil, []) + | p1 = tac2pat; "::"; p2 = tac2pat -> + CPatRef (!@loc, AbsKn Tac2core.Core.c_cons, [p1; p2]) + ] | "0" [ "_" -> CPatAny (!@loc) | "()" -> CPatTup (Loc.tag ~loc:!@loc []) -- cgit v1.2.3 From fbfe239730bd5069026ae4e5356e93d3f3bfcb53 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 24 Jul 2017 19:37:03 +0200 Subject: Correctly pushing variables for tuple patterns. --- src/tac2intern.ml | 8 +++++--- 1 file changed, 5 insertions(+), 3 deletions(-) diff --git a/src/tac2intern.ml b/src/tac2intern.ml index b63e6a0cd8..ffbdaf4b9b 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -817,11 +817,13 @@ and intern_case env loc e pl = end | p -> todo ~loc:(loc_of_patexpr p) () in - let ids = Array.map_of_list map pl in - let tc = GTypTuple (List.map (fun _ -> GTypVar (fresh_id env)) pl) in + let ids = List.map map pl in + let targs = List.map (fun _ -> GTypVar (fresh_id env)) pl in + let tc = GTypTuple targs in let () = unify ~loc:(loc_of_tacexpr e) env t tc in + let env = List.fold_left2 (fun env na t -> push_name na (monomorphic t) env) env ids targs in let (b, tb) = intern_rec env b in - (GTacCse (e', GCaseTuple len, [||], [|ids, b|]), tb) + (GTacCse (e', GCaseTuple len, [||], [|Array.of_list ids, b|]), tb) | (p, _) :: _ -> todo ~loc:(loc_of_patexpr p) () end | PKind_variant kn -> -- cgit v1.2.3 From 41cea8603b35a1af405650d8a2b9aaa89a445367 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 24 Jul 2017 19:00:29 +0200 Subject: Adding a few primitive functions. --- src/tac2core.ml | 197 +++++++++++++++++++++++++++++++++++++++++++++++++++-- src/tac2core.mli | 9 +++ tests/ltac2.v | 12 ++++ theories/Constr.v | 16 ++--- theories/Control.v | 7 ++ theories/Init.v | 2 + theories/Message.v | 2 + 7 files changed, 233 insertions(+), 12 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 91a3bfa168..13aa44c815 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -27,6 +27,15 @@ let val_tag t = match val_tag t with let val_constr = val_tag (topwit Stdarg.wit_constr) let val_ident = val_tag (topwit Stdarg.wit_ident) let val_pp = Val.create "ltac2:pp" +let val_sort = Val.create "ltac2:sort" +let val_cast = Val.create "ltac2:cast" +let val_inductive = Val.create "ltac2:inductive" +let val_constant = Val.create "ltac2:constant" +let val_constructor = Val.create "ltac2:constructor" +let val_projection = Val.create "ltac2:projection" +let val_univ = Val.create "ltac2:universe" +let val_kont : (Exninfo.iexn -> valexpr Proofview.tactic) Val.typ = + Val.create "ltac2:kont" let extract_val (type a) (tag : a Val.typ) (Val.Dyn (tag', v)) : a = match Val.eq tag tag' with @@ -121,9 +130,39 @@ let to_exn c = match c with | ValOpn (kn, c) -> (LtacError (kn, c), Exninfo.null) | _ -> to_ext val_exn c +let of_option = function +| None -> ValInt 0 +| Some c -> ValBlk (0, [|c|]) + +let to_option = function +| ValInt 0 -> None +| ValBlk (0, [|c|]) -> Some c +| _ -> assert false + let of_pp c = of_ext val_pp c let to_pp c = to_ext val_pp c +let of_tuple cl = ValBlk (0, cl) +let to_tuple = function +| ValBlk (0, cl) -> cl +| _ -> assert false + +let of_array = of_tuple +let to_array = to_tuple + +let of_name c = match c with +| Anonymous -> of_option None +| Name id -> of_option (Some (of_ident id)) + +let of_instance sigma u = + let u = Univ.Instance.to_array (EConstr.EInstance.kind sigma u) in + of_array (Array.map (fun v -> of_ext val_univ v) u) + +let of_rec_declaration (nas, ts, cs) = + (of_array (Array.map of_name nas), + of_array (Array.map of_constr ts), + of_array (Array.map of_constr cs)) + end let val_valexpr = Val.create "ltac2:valexpr" @@ -174,11 +213,11 @@ let prm_print : ml_tactic = function | _ -> assert false let prm_message_of_int : ml_tactic = function -| [ValInt s] -> return (ValExt (Val.Dyn (val_pp, int s))) +| [ValInt s] -> return (Value.of_pp (int s)) | _ -> assert false let prm_message_of_string : ml_tactic = function -| [ValStr s] -> return (ValExt (Val.Dyn (val_pp, str (Bytes.to_string s)))) +| [ValStr s] -> return (Value.of_pp (str (Bytes.to_string s))) | _ -> assert false let prm_message_of_constr : ml_tactic = function @@ -186,10 +225,17 @@ let prm_message_of_constr : ml_tactic = function pf_apply begin fun env sigma -> let c = Value.to_constr c in let pp = Printer.pr_econstr_env env sigma c in - return (ValExt (Val.Dyn (val_pp, pp))) + return (Value.of_pp pp) end | _ -> assert false +let prm_message_of_ident : ml_tactic = function +| [c] -> + let c = Value.to_ident c in + let pp = Id.print c in + return (Value.of_pp pp) +| _ -> assert false + let prm_message_concat : ml_tactic = function | [m1; m2] -> let m1 = Value.to_pp m1 in @@ -298,6 +344,101 @@ let prm_constr_equal : ml_tactic = function Proofview.tclUNIT (Value.of_bool b) | _ -> assert false +let prm_constr_kind : ml_tactic = function +| [c] -> + let open Constr in + Proofview.tclEVARMAP >>= fun sigma -> + let c = Value.to_constr c in + return begin match EConstr.kind sigma c with + | Rel n -> + ValBlk (0, [|Value.of_int n|]) + | Var id -> + ValBlk (1, [|Value.of_ident id|]) + | Meta n -> + ValBlk (2, [|Value.of_int n|]) + | Evar (evk, args) -> + ValBlk (3, [| + Value.of_int (Evar.repr evk); + Value.of_array (Array.map Value.of_constr args) + |]) + | Sort s -> + ValBlk (4, [|Value.of_ext val_sort s|]) + | Cast (c, k, t) -> + ValBlk (5, [| + Value.of_constr c; + Value.of_ext val_cast k; + Value.of_constr t; + |]) + | Prod (na, t, u) -> + ValBlk (6, [| + Value.of_name na; + Value.of_constr t; + Value.of_constr u; + |]) + | Lambda (na, t, c) -> + ValBlk (7, [| + Value.of_name na; + Value.of_constr t; + Value.of_constr c; + |]) + | LetIn (na, b, t, c) -> + ValBlk (8, [| + Value.of_name na; + Value.of_constr b; + Value.of_constr t; + Value.of_constr c; + |]) + | App (c, cl) -> + ValBlk (9, [| + Value.of_constr c; + Value.of_array (Array.map Value.of_constr cl) + |]) + | Const (cst, u) -> + ValBlk (10, [| + Value.of_ext val_constant cst; + Value.of_instance sigma u; + |]) + | Ind (ind, u) -> + ValBlk (11, [| + Value.of_ext val_inductive ind; + Value.of_instance sigma u; + |]) + | Construct (cstr, u) -> + ValBlk (12, [| + Value.of_ext val_constructor cstr; + Value.of_instance sigma u; + |]) + | Case (_, c, t, bl) -> + ValBlk (13, [| + Value.of_constr c; + Value.of_constr t; + Value.of_array (Array.map (fun c -> Value.of_constr c) bl); + |]) + | Fix ((recs, i), def) -> + let (nas, ts, cs) = Value.of_rec_declaration def in + ValBlk (14, [| + Value.of_array (Array.map Value.of_int recs); + Value.of_int i; + nas; + ts; + cs; + |]) + | CoFix (i, def) -> + let (nas, ts, cs) = Value.of_rec_declaration def in + ValBlk (15, [| + Value.of_int i; + nas; + ts; + cs; + |]) + | Proj (p, c) -> + ValBlk (16, [| + Value.of_ext val_projection p; + Value.of_constr c; + |]) + end +| _ -> assert false + (** Error *) let prm_throw : ml_tactic = function @@ -315,7 +456,7 @@ let prm_zero : ml_tactic = function Proofview.tclZERO ~info e | _ -> assert false -(** exn -> 'a *) +(** (unit -> 'a) -> (exn -> 'a) -> 'a *) let prm_plus : ml_tactic = function | [x; k] -> Proofview.tclOR (thaw x) (fun e -> interp_app k [Value.of_exn e]) @@ -352,6 +493,30 @@ let prm_enter : ml_tactic = function Proofview.tclINDEPENDENT f >>= fun () -> return v_unit | _ -> assert false +let k_var = Id.of_string "k" +let e_var = Id.of_string "e" +let prm_apply_kont_h = pname "apply_kont" + +(** (unit -> 'a) -> ('a * ('exn -> 'a)) result *) +let prm_case : ml_tactic = function +| [f] -> + Proofview.tclCASE (thaw f) >>= begin function + | Proofview.Next (x, k) -> + let k = { + clos_env = Id.Map.singleton k_var (Value.of_ext val_kont k); + clos_var = [Name e_var]; + clos_exp = GTacPrm (prm_apply_kont_h, [GTacVar k_var; GTacVar e_var]); + } in + return (ValBlk (0, [| Value.of_tuple [| x; ValCls k |] |])) + | Proofview.Fail e -> return (ValBlk (1, [| Value.of_exn e |])) + end +| _ -> assert false + +(** 'a kont -> exn -> 'a *) +let prm_apply_kont : ml_tactic = function +| [k; e] -> (Value.to_ext val_kont k) (Value.to_exn e) +| _ -> assert false + (** int -> int -> (unit -> 'a) -> 'a *) let prm_focus : ml_tactic = function | [i; j; tac] -> @@ -400,6 +565,25 @@ let prm_hyp : ml_tactic = function end | _ -> assert false +let prm_hyps : ml_tactic = function +| [_] -> + pf_apply begin fun env _ -> + let open Context.Named.Declaration in + let hyps = Environ.named_context env in + let map = function + | LocalAssum (id, t) -> + let t = EConstr.of_constr t in + Value.of_tuple [|Value.of_ident id; Value.of_option None; Value.of_constr t|] + | LocalDef (id, c, t) -> + let c = EConstr.of_constr c in + let t = EConstr.of_constr t in + Value.of_tuple [|Value.of_ident id; Value.of_option (Some (Value.of_constr c)); Value.of_constr t|] + in + let hyps = List.rev_map map hyps in + return (Value.of_list hyps) + end +| _ -> assert false + (** (unit -> constr) -> unit *) let prm_refine : ml_tactic = function | [c] -> @@ -416,6 +600,7 @@ let () = Tac2env.define_primitive (pname "print") prm_print let () = Tac2env.define_primitive (pname "message_of_string") prm_message_of_string let () = Tac2env.define_primitive (pname "message_of_int") prm_message_of_int let () = Tac2env.define_primitive (pname "message_of_constr") prm_message_of_constr +let () = Tac2env.define_primitive (pname "message_of_ident") prm_message_of_ident let () = Tac2env.define_primitive (pname "message_concat") prm_message_concat let () = Tac2env.define_primitive (pname "array_make") prm_array_make @@ -430,6 +615,7 @@ let () = Tac2env.define_primitive (pname "string_set") prm_string_set let () = Tac2env.define_primitive (pname "constr_type") prm_constr_type let () = Tac2env.define_primitive (pname "constr_equal") prm_constr_equal +let () = Tac2env.define_primitive (pname "constr_kind") prm_constr_kind let () = Tac2env.define_primitive (pname "int_equal") prm_int_equal let () = Tac2env.define_primitive (pname "int_compare") prm_int_compare @@ -446,6 +632,8 @@ let () = Tac2env.define_primitive (pname "once") prm_once let () = Tac2env.define_primitive (pname "dispatch") prm_dispatch let () = Tac2env.define_primitive (pname "extend") prm_extend let () = Tac2env.define_primitive (pname "enter") prm_enter +let () = Tac2env.define_primitive (pname "case") prm_case +let () = Tac2env.define_primitive (pname "apply_kont") prm_apply_kont let () = Tac2env.define_primitive (pname "focus") prm_focus let () = Tac2env.define_primitive (pname "shelve") prm_shelve @@ -453,6 +641,7 @@ let () = Tac2env.define_primitive (pname "shelve_unifiable") prm_shelve_unifiabl let () = Tac2env.define_primitive (pname "new_goal") prm_new_goal let () = Tac2env.define_primitive (pname "goal") prm_goal let () = Tac2env.define_primitive (pname "hyp") prm_hyp +let () = Tac2env.define_primitive (pname "hyps") prm_hyps let () = Tac2env.define_primitive (pname "refine") prm_refine (** ML types *) diff --git a/src/tac2core.mli b/src/tac2core.mli index fc90499ac6..41c79b2c65 100644 --- a/src/tac2core.mli +++ b/src/tac2core.mli @@ -59,4 +59,13 @@ val to_exn : valexpr -> Exninfo.iexn val of_ident : Id.t -> valexpr val to_ident : valexpr -> Id.t +val of_array : valexpr array -> valexpr +val to_array : valexpr -> valexpr array + +val of_tuple : valexpr array -> valexpr +val to_tuple : valexpr -> valexpr array + +val of_option : valexpr option -> valexpr +val to_option : valexpr -> valexpr option + end diff --git a/tests/ltac2.v b/tests/ltac2.v index 3d3d0032c5..770d385406 100644 --- a/tests/ltac2.v +++ b/tests/ltac2.v @@ -117,3 +117,15 @@ Goal True. Proof. refine! I. Abort. + +Goal True. +Proof. +let x _ := plus (fun _ => 0) (fun _ => 1) in +match case x with +| Val x => + match x with + | (x, k) => Message.print (Message.of_int (k Not_found)) + end +| Err x => Message.print (Message.of_string "Err") +end. +Abort. diff --git a/theories/Constr.v b/theories/Constr.v index c340e3aa87..d7cd3b58a3 100644 --- a/theories/Constr.v +++ b/theories/Constr.v @@ -16,28 +16,28 @@ Ltac2 @ external equal : constr -> constr -> bool := "ltac2" "constr_equal". Module Unsafe. -(** Low-level access to kernel term. Use with care! *) +(** Low-level access to kernel terms. Use with care! *) Ltac2 Type kind := [ | Rel (int) | Var (ident) | Meta (meta) -| Evar (evar, constr list) +| Evar (evar, constr array) | Sort (sort) | Cast (constr, cast, constr) | Prod (ident option, constr, constr) | Lambda (ident option, constr, constr) | LetIn (ident option, constr, constr, constr) -| App (constr, constr list) +| App (constr, constr array) | Constant (constant, instance) | Ind (inductive, instance) | Constructor (inductive, instance) -(* - | Case of case_info * 'constr * 'constr * 'constr array - | Fix of ('constr, 'types) pfixpoint - | CoFix of ('constr, 'types) pcofixpoint -*) +| Case (constr, constr, constr array) +| Fix (int array, int, ident option array, constr array, constr array) +| CoFix (int, ident option array, constr array, constr array) | Proj (projection, constr) ]. +Ltac2 @ external kind : constr -> kind := "ltac2" "constr_kind". + End Unsafe. diff --git a/theories/Control.v b/theories/Control.v index 3bc572547c..a6d46a89a8 100644 --- a/theories/Control.v +++ b/theories/Control.v @@ -21,6 +21,7 @@ Ltac2 @ external once : (unit -> 'a) -> 'a := "ltac2" "once". Ltac2 @ external dispatch : (unit -> unit) list -> unit := "ltac2" "dispatch". Ltac2 @ external extend : (unit -> unit) list -> (unit -> unit) -> (unit -> unit) list -> unit := "ltac2" "extend". Ltac2 @ external enter : (unit -> unit) -> unit := "ltac2" "enter". +Ltac2 @ external case : (unit -> 'a) -> ('a * (exn -> 'a)) result := "ltac2" "case". (** Proof state manipulation *) @@ -44,6 +45,12 @@ Ltac2 @ external hyp : ident -> constr := "ltac2" "hyp". goal under focus, looks for the section variable with the given name. If there is one, looks for the hypothesis with the given name. *) +Ltac2 @ external hyps : unit -> (ident * constr option * constr) list := "ltac2" "hyps". +(** Panics if there is more than one goal under focus. If there is no + goal under focus, returns the list of section variables. + If there is one, returns the list of hypotheses. In both cases, the + list is ordered with rightmost values being last introduced. *) + (** Refinement *) Ltac2 @ external refine : (unit -> constr) -> unit := "ltac2" "refine". diff --git a/theories/Init.v b/theories/Init.v index 1d2d40f5c0..c0a73576d3 100644 --- a/theories/Init.v +++ b/theories/Init.v @@ -41,6 +41,8 @@ Ltac2 Type 'a ref := { mutable contents : 'a }. Ltac2 Type bool := [ true | false ]. +Ltac2 Type 'a result := [ Val ('a) | Err (exn) ]. + (** Pervasive exceptions *) Ltac2 Type exn ::= [ Out_of_bounds ]. diff --git a/theories/Message.v b/theories/Message.v index b2159612cb..45f4b221db 100644 --- a/theories/Message.v +++ b/theories/Message.v @@ -14,6 +14,8 @@ Ltac2 @ external of_string : string -> message := "ltac2" "message_of_string". Ltac2 @ external of_int : int -> message := "ltac2" "message_of_int". +Ltac2 @ external of_ident : ident -> message := "ltac2" "message_of_ident". + Ltac2 @ external of_constr : constr -> message := "ltac2" "message_of_constr". (** Panics if there is more than one goal under focus. *) -- cgit v1.2.3 From 5748cd3a913eec7a24600715fc9b71044a7c38b1 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 25 Jul 2017 12:27:31 +0200 Subject: Generalizing patterns in fun bindings. --- src/g_ltac2.ml4 | 4 ++-- src/tac2core.ml | 2 +- src/tac2expr.mli | 2 +- src/tac2intern.ml | 44 ++++++++++++++++++++++++++++++++++++++------ 4 files changed, 42 insertions(+), 10 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 6cdbccb11d..47def14125 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -125,8 +125,8 @@ GEXTEND Gram | l = Prim.ident -> Loc.tag ~loc:!@loc (Name l) ] ] ; input_fun: - [ [ b = binder -> (b, None) - | "("; b = binder; ":"; t = tac2type; ")" -> (b, Some t) ] ] + [ [ b = tac2pat LEVEL "0" -> (b, None) + | "("; b = tac2pat; t = OPT [ ":"; t = tac2type -> t ]; ")" -> (b, t) ] ] ; tac2def_body: [ [ name = binder; it = LIST0 input_fun; ":="; e = tac2expr -> diff --git a/src/tac2core.ml b/src/tac2core.ml index 13aa44c815..b665f761ce 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -734,7 +734,7 @@ let dummy_loc = Loc.make_loc (-1, -1) let rthunk e = let loc = Tac2intern.loc_of_tacexpr e in - let var = [Loc.tag ~loc Anonymous, Some (CTypRef (loc, AbsKn Core.t_unit, []))] in + let var = [CPatAny loc, Some (CTypRef (loc, AbsKn Core.t_unit, []))] in CTacFun (loc, var, e) let add_generic_scope s entry arg = diff --git a/src/tac2expr.mli b/src/tac2expr.mli index acdad9bab4..a9f2109cb2 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -83,7 +83,7 @@ type raw_patexpr = type raw_tacexpr = | CTacAtm of atom located | CTacRef of tacref or_relid -| CTacFun of Loc.t * (Name.t located * raw_typexpr option) list * raw_tacexpr +| CTacFun of Loc.t * (raw_patexpr * raw_typexpr option) list * raw_tacexpr | CTacApp of Loc.t * raw_tacexpr * raw_tacexpr list | CTacLet of Loc.t * rec_flag * (Name.t located * raw_typexpr option * raw_tacexpr) list * raw_tacexpr | CTacTup of raw_tacexpr list located diff --git a/src/tac2intern.ml b/src/tac2intern.ml index ffbdaf4b9b..b0ba9adf5e 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -585,6 +585,14 @@ let is_constructor env qid = match get_variable env qid with | ArgArg (TacConstructor _) -> true | _ -> false +(** Used to generate a fresh tactic variable for pattern-expansion *) +let fresh_var env = + let bad id = + Id.Map.mem id env.env_var || + (try ignore (locate_ltac (qualid_of_ident id)); true with Not_found -> false) + in + Namegen.next_ident_away_from (Id.of_string "p") bad + let rec intern_rec env = function | CTacAtm (_, atm) -> intern_atm env atm | CTacRef qid as e -> @@ -600,16 +608,24 @@ let rec intern_rec env = function intern_constructor env loc kn [] end | CTacFun (loc, bnd, e) -> - let fold (env, bnd, tl) ((_, na), t) = + let fold (env, bnd, tl) (pat, t) = let t = match t with | None -> GTypVar (fresh_id env) | Some t -> intern_type env t in - let env = push_name na (monomorphic t) env in - (env, na :: bnd, t :: tl) + let id = fresh_var env in + let env = push_name (Name id) (monomorphic t) env in + (env, (id, pat) :: bnd, t :: tl) in let (env, bnd, tl) = List.fold_left fold (env, [], []) bnd in - let bnd = List.rev bnd in + (** Expand pattern: [fun p => t] becomes [fun x => match x with p => t end] *) + let fold e (id, pat) = + let loc = loc_of_patexpr pat in + let qid = RelId (Loc.tag ~loc (qualid_of_ident id)) in + CTacCse (loc, CTacRef qid, [pat, e]) + in + let e = List.fold_left fold e bnd in + let bnd = List.rev_map (fun (id, _) -> Name id) bnd in let (e, t) = intern_rec env e in let t = List.fold_left (fun accu t -> GTypArrow (t, accu)) t tl in (GTacFun (bnd, e), t) @@ -1125,6 +1141,17 @@ let get_projection0 var = match var with kn | AbsKn kn -> kn +let rec ids_of_pattern accu = function +| CPatAny _ -> accu +| CPatRef (_, RelId (_, qid), pl) -> + let (dp, id) = repr_qualid qid in + let accu = if DirPath.is_empty dp then Id.Set.add id accu else accu in + List.fold_left ids_of_pattern accu pl +| CPatRef (_, AbsKn _, pl) -> + List.fold_left ids_of_pattern accu pl +| CPatTup (_, pl) -> + List.fold_left ids_of_pattern accu pl + let rec globalize ids e = match e with | CTacAtm _ -> e | CTacRef ref -> @@ -1134,8 +1161,13 @@ let rec globalize ids e = match e with | ArgArg kn -> CTacRef (AbsKn kn) end | CTacFun (loc, bnd, e) -> - let fold accu ((_, na), _) = add_name accu na in - let ids = List.fold_left fold ids bnd in + let fold (pats, accu) (pat, t) = + let accu = ids_of_pattern accu pat in + let pat = globalize_pattern ids pat in + ((pat, t) :: pats, accu) + in + let bnd, ids = List.fold_left fold ([], ids) bnd in + let bnd = List.rev bnd in let e = globalize ids e in CTacFun (loc, bnd, e) | CTacApp (loc, e, el) -> -- cgit v1.2.3 From a3da80680400610ffe8d7de33d9ca1ee1106ae28 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 26 Jul 2017 00:32:50 +0200 Subject: Bugfix: wrong access to non-constant constructor compilation. --- src/tac2intern.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/tac2intern.ml b/src/tac2intern.ml index b0ba9adf5e..06f04c4c3d 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -874,7 +874,7 @@ and intern_case env loc e pl = (succ ncst, narg) else let () = - if Option.is_empty const.(narg) then + if Option.is_empty nonconst.(narg) then let ids = Array.map_of_list (fun _ -> Anonymous) args in nonconst.(narg) <- Some (ids, br') in -- cgit v1.2.3 From bbf4ee2fe5072fa0bb639dce649c16fdd76f44b0 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 25 Jul 2017 23:43:26 +0200 Subject: Exporting some basic tactics from Ltac1. --- _CoqProject | 3 + src/ltac2_plugin.mlpack | 1 + src/tac2stdlib.ml | 166 ++++++++++++++++++++++++++++++++++++++++++++++++ src/tac2stdlib.mli | 9 +++ theories/Ltac2.v | 1 + theories/Std.v | 59 +++++++++++++++++ 6 files changed, 239 insertions(+) create mode 100644 src/tac2stdlib.ml create mode 100644 src/tac2stdlib.mli create mode 100644 theories/Std.v diff --git a/_CoqProject b/_CoqProject index 90338abbfb..6d3470cfa7 100644 --- a/_CoqProject +++ b/_CoqProject @@ -15,6 +15,8 @@ src/tac2entries.ml src/tac2entries.mli src/tac2core.ml src/tac2core.mli +src/tac2stdlib.ml +src/tac2stdlib.mli src/g_ltac2.ml4 src/ltac2_plugin.mlpack @@ -25,4 +27,5 @@ theories/Array.v theories/Control.v theories/Message.v theories/Constr.v +theories/Std.v theories/Ltac2.v diff --git a/src/ltac2_plugin.mlpack b/src/ltac2_plugin.mlpack index 3d87a8cddb..dc78207291 100644 --- a/src/ltac2_plugin.mlpack +++ b/src/ltac2_plugin.mlpack @@ -4,4 +4,5 @@ Tac2intern Tac2interp Tac2entries Tac2core +Tac2stdlib G_ltac2 diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml new file mode 100644 index 0000000000..89a8d98693 --- /dev/null +++ b/src/tac2stdlib.ml @@ -0,0 +1,166 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* return v_unit + +let wrap f = + return () >>= fun () -> return (f ()) + +let wrap_unit f = + return () >>= fun () -> f (); return v_unit + +let define_prim0 name tac = + let tac = function + | [_] -> lift tac + | _ -> assert false + in + Tac2env.define_primitive (pname name) tac + +let define_prim1 name tac = + let tac = function + | [x] -> lift (tac x) + | _ -> assert false + in + Tac2env.define_primitive (pname name) tac + +let define_prim2 name tac = + let tac = function + | [x; y] -> lift (tac x y) + | _ -> assert false + in + Tac2env.define_primitive (pname name) tac + +(** Tactics from coretactics *) + +let () = define_prim0 "tac_reflexivity" Tactics.intros_reflexivity + +(* + +TACTIC EXTEND exact + [ "exact" casted_constr(c) ] -> [ Tactics.exact_no_check c ] +END + +*) + +let () = define_prim0 "tac_assumption" Tactics.assumption + +let () = define_prim1 "tac_transitivity" begin fun c -> + let c = Value.to_constr c in + Tactics.intros_transitivity (Some c) +end + +let () = define_prim0 "tac_etransitivity" (Tactics.intros_transitivity None) + +let () = define_prim1 "tac_cut" begin fun c -> + let c = Value.to_constr c in + Tactics.cut c +end + +let () = define_prim0 "tac_left" (Tactics.left_with_bindings false NoBindings) +let () = define_prim0 "tac_eleft" (Tactics.left_with_bindings true NoBindings) +let () = define_prim0 "tac_right" (Tactics.right_with_bindings false NoBindings) +let () = define_prim0 "tac_eright" (Tactics.right_with_bindings true NoBindings) + +let () = define_prim1 "tac_exactnocheck" begin fun c -> + Tactics.exact_no_check (Value.to_constr c) +end + +let () = define_prim1 "tac_vmcastnocheck" begin fun c -> + Tactics.vm_cast_no_check (Value.to_constr c) +end + +let () = define_prim1 "tac_nativecastnocheck" begin fun c -> + Tactics.native_cast_no_check (Value.to_constr c) +end + +let () = define_prim0 "tac_constructor" (Tactics.any_constructor false None) +let () = define_prim0 "tac_econstructor" (Tactics.any_constructor true None) + +let () = define_prim1 "tac_constructorn" begin fun n -> + let n = Value.to_int n in + Tactics.constructor_tac false None n NoBindings +end + +let () = define_prim1 "tac_econstructorn" begin fun n -> + let n = Value.to_int n in + Tactics.constructor_tac true None n NoBindings +end + +let () = define_prim0 "tac_symmetry" (Tactics.intros_symmetry Locusops.onConcl) + +let () = define_prim0 "tac_split" (Tactics.split_with_bindings false [NoBindings]) +let () = define_prim0 "tac_esplit" (Tactics.split_with_bindings true [NoBindings]) + +let () = define_prim1 "tac_rename" begin fun ids -> + let ids = Value.to_list ids in + let map c = match Value.to_tuple c with + | [|x; y|] -> (Value.to_ident x, Value.to_ident y) + | _ -> assert false + in + let ids = List.map map ids in + Tactics.rename_hyp ids +end + +let () = define_prim1 "tac_revert" begin fun ids -> + let ids = List.map Value.to_ident (Value.to_list ids) in + Tactics.revert ids +end + +let () = define_prim0 "tac_admit" Proofview.give_up + +let () = define_prim2 "tac_fix" begin fun idopt n -> + let idopt = Option.map Value.to_ident (Value.to_option idopt) in + let n = Value.to_int n in + Tactics.fix idopt n +end + +let () = define_prim1 "tac_cofix" begin fun idopt -> + let idopt = Option.map Value.to_ident (Value.to_option idopt) in + Tactics.cofix idopt +end + +let () = define_prim1 "tac_clear" begin fun ids -> + let ids = List.map Value.to_ident (Value.to_list ids) in + Tactics.clear ids +end + +let () = define_prim1 "tac_keep" begin fun ids -> + let ids = List.map Value.to_ident (Value.to_list ids) in + Tactics.keep ids +end + +let () = define_prim1 "tac_clearbody" begin fun ids -> + let ids = List.map Value.to_ident (Value.to_list ids) in + Tactics.clear_body ids +end + +(** Tactics from extratactics *) + +let () = define_prim1 "tac_absurd" begin fun c -> + Contradiction.absurd (Value.to_constr c) +end + +let () = define_prim1 "tac_subst" begin fun ids -> + let ids = List.map Value.to_ident (Value.to_list ids) in + Equality.subst ids +end + +let () = define_prim0 "tac_substall" (return () >>= fun () -> Equality.subst_all ()) diff --git a/src/tac2stdlib.mli b/src/tac2stdlib.mli new file mode 100644 index 0000000000..927b57074d --- /dev/null +++ b/src/tac2stdlib.mli @@ -0,0 +1,9 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* unit := "ltac2" "tac_reflexivity". + +Ltac2 @ external assumption : unit -> unit := "ltac2" "tac_assumption". + +Ltac2 @ external transitivity : constr -> unit := "ltac2" "tac_transitivity". + +Ltac2 @ external etransitivity : unit -> unit := "ltac2" "tac_etransitivity". + +Ltac2 @ external cut : constr -> unit := "ltac2" "tac_cut". + +Ltac2 @ external left : unit -> unit := "ltac2" "tac_left". +Ltac2 @ external eleft : unit -> unit := "ltac2" "tac_eleft". +Ltac2 @ external right : unit -> unit := "ltac2" "tac_right". +Ltac2 @ external eright : unit -> unit := "ltac2" "tac_eright". + +Ltac2 @ external constructor : unit -> unit := "ltac2" "tac_constructor". +Ltac2 @ external econstructor : unit -> unit := "ltac2" "tac_econstructor". +Ltac2 @ external split : unit -> unit := "ltac2" "tac_split". +Ltac2 @ external esplit : unit -> unit := "ltac2" "tac_esplit". + +Ltac2 @ external constructor_n : int -> unit := "ltac2" "tac_constructorn". +Ltac2 @ external econstructor_n : int -> unit := "ltac2" "tac_econstructorn". + +Ltac2 @ external symmetry : unit -> unit := "ltac2" "tac_symmetry". + +Ltac2 @ external rename : (ident * ident) list -> unit := "ltac2" "tac_rename". + +Ltac2 @ external revert : ident list -> unit := "ltac2" "tac_revert". + +Ltac2 @ external admit : unit -> unit := "ltac2" "tac_admit". + +Ltac2 @ external fix_ : ident option -> int -> unit := "ltac2" "tac_fix". +Ltac2 @ external cofix_ : ident option -> unit := "ltac2" "tac_cofix". + +Ltac2 @ external clear : ident list -> unit := "ltac2" "tac_clear". +Ltac2 @ external keep : ident list -> unit := "ltac2" "tac_keep". + +Ltac2 @ external clearbody : ident list -> unit := "ltac2" "tac_clearbody". + +Ltac2 @ external exact_no_check : constr -> unit := "ltac2" "tac_exactnocheck". +Ltac2 @ external vm_cast_no_check : constr -> unit := "ltac2" "tac_vmcastnocheck". +Ltac2 @ external native_cast_no_check : constr -> unit := "ltac2" "tac_nativecastnocheck". + +Ltac2 @ external absurd : constr -> unit := "ltac2" "tac_absurd". + +Ltac2 @ external subst : ident list -> unit := "ltac2" "tac_subst". +Ltac2 @ external subst_all : unit -> unit := "ltac2" "tac_substall". -- cgit v1.2.3 From 93b5d42467dae8513ad7da4990f909bcc9f5b7fa Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 26 Jul 2017 13:30:58 +0200 Subject: Properly handling toplevel recursive definitions. --- src/tac2entries.ml | 124 +++++++++++++++++++++++++++++++++-------------------- src/tac2intern.mli | 1 + 2 files changed, 79 insertions(+), 46 deletions(-) diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 46f390a6d4..100041f15e 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -241,56 +241,88 @@ let inTypExt : typext -> obj = let dummy_loc = Loc.make_loc (-1, -1) +let fresh_var avoid x = + let bad id = + Id.Set.mem id avoid || + (try ignore (Tac2env.locate_ltac (qualid_of_ident id)); true with Not_found -> false) + in + Namegen.next_ident_away_from (Id.of_string x) bad + +(** Mangle recursive tactics *) +let inline_rec_tactic tactics = + let avoid = List.fold_left (fun accu ((_, id), _) -> Id.Set.add id accu) Id.Set.empty tactics in + let map (id, e) = match e with + | CTacFun (loc, pat, _) -> (id, pat, e) + | _ -> + let loc, _ = id in + user_err ?loc (str "Recursive tactic definitions must be functions") + in + let tactics = List.map map tactics in + let map (id, pat, e) = + let fold_var (avoid, ans) (pat, _) = + let id = fresh_var avoid "x" in + let loc = loc_of_patexpr pat in + (Id.Set.add id avoid, Loc.tag ~loc id :: ans) + in + (** Fresh variables to abstract over the function patterns *) + let _, vars = List.fold_left fold_var (avoid, []) pat in + let map_body ((loc, id), _, e) = (loc, Name id), None, e in + let bnd = List.map map_body tactics in + let pat_of_id (loc, id) = + let qid = (loc, qualid_of_ident id) in + (CPatRef (Option.default dummy_loc loc, RelId qid, []), None) + in + let var_of_id (loc, id) = + let qid = (loc, qualid_of_ident id) in + CTacRef (RelId qid) + in + let loc0 = loc_of_tacexpr e in + let vpat = List.map pat_of_id vars in + let varg = List.map var_of_id vars in + let e = CTacLet (loc0, true, bnd, CTacApp (loc0, var_of_id id, varg)) in + (id, CTacFun (loc0, vpat, e)) + in + List.map map tactics + let register_ltac ?(local = false) isrec tactics = - if isrec then - let map (na, e) = (na, None, e) in - let bindings = List.map map tactics in - let map ((loc, na), e) = match na with - | Anonymous -> None - | Name id -> - let qid = Libnames.qualid_of_ident id in - let e = CTacLet (dummy_loc, true, bindings, CTacRef (RelId (loc, qid))) in - let (e, t) = intern e in - let e = match e with - | GTacLet (true, _, e) -> assert false - | _ -> assert false - in - Some (e, t) + let map ((loc, na), e) = + let id = match na with + | Anonymous -> + user_err ?loc (str "Tactic definition must have a name") + | Name id -> id in - let tactics = List.map map tactics in - assert false (** FIXME *) - else - let map ((loc, na), e) = - let (e, t) = intern e in - let () = - if not (is_value e) then - user_err ?loc (str "Tactic definition must be a syntactical value") - in - let id = match na with - | Anonymous -> - user_err ?loc (str "Tactic definition must have a name") - | Name id -> id - in - let kn = Lib.make_kn id in - let exists = - try let _ = Tac2env.interp_global kn in true with Not_found -> false - in - let () = - if exists then - user_err ?loc (str "Tactic " ++ Nameops.pr_id id ++ str " already exists") - in - (id, e, t) + ((loc, id), e) + in + let tactics = List.map map tactics in + let tactics = + if isrec then inline_rec_tactic tactics else tactics + in + let map ((loc, id), e) = + let (e, t) = intern e in + let () = + if not (is_value e) then + user_err ?loc (str "Tactic definition must be a syntactical value") in - let defs = List.map map tactics in - let iter (id, e, t) = - let def = { - tacdef_local = local; - tacdef_expr = e; - tacdef_type = t; - } in - ignore (Lib.add_leaf id (inTacDef def)) + let kn = Lib.make_kn id in + let exists = + try let _ = Tac2env.interp_global kn in true with Not_found -> false in - List.iter iter defs + let () = + if exists then + user_err ?loc (str "Tactic " ++ Nameops.pr_id id ++ str " already exists") + in + (id, e, t) + in + let defs = List.map map tactics in + let iter (id, e, t) = + let def = { + tacdef_local = local; + tacdef_expr = e; + tacdef_type = t; + } in + ignore (Lib.add_leaf id (inTacDef def)) + in + List.iter iter defs let qualid_to_ident (loc, qid) = let (dp, id) = Libnames.repr_qualid qid in diff --git a/src/tac2intern.mli b/src/tac2intern.mli index 3d400a5cdd..b2604c4ea7 100644 --- a/src/tac2intern.mli +++ b/src/tac2intern.mli @@ -12,6 +12,7 @@ open Mod_subst open Tac2expr val loc_of_tacexpr : raw_tacexpr -> Loc.t +val loc_of_patexpr : raw_patexpr -> Loc.t val intern : raw_tacexpr -> glb_tacexpr * type_scheme val intern_typedef : (KerName.t * int) Id.Map.t -> raw_quant_typedef -> glb_quant_typedef -- cgit v1.2.3 From edb9a5dd11520804d65e9b1e95ca38bb4acbb0e6 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 26 Jul 2017 16:01:52 +0200 Subject: Lightweight quotation syntax for terms and idents. --- src/g_ltac2.ml4 | 2 ++ 1 file changed, 2 insertions(+) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 47def14125..605cb75d66 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -91,6 +91,8 @@ GEXTEND Gram [ [ n = Prim.integer -> CTacAtm (Loc.tag ~loc:!@loc (AtmInt n)) | s = Prim.string -> CTacAtm (Loc.tag ~loc:!@loc (AtmStr s)) | id = Prim.qualid -> CTacRef (RelId id) + | "@"; id = Prim.ident -> inj_ident !@loc id + | "'"; c = Constr.constr -> inj_open_constr !@loc c | IDENT "constr"; ":"; "("; c = Constr.lconstr; ")" -> inj_constr !@loc c | IDENT "open_constr"; ":"; "("; c = Constr.lconstr; ")" -> inj_open_constr !@loc c | IDENT "ident"; ":"; "("; c = Prim.ident; ")" -> inj_ident !@loc c -- cgit v1.2.3 From b13693a39014d727787c003c6d445c3bb6f2aef6 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 26 Jul 2017 16:11:32 +0200 Subject: Better typing errors for function types. --- src/tac2intern.ml | 31 +++++++++++++++++++++++++++---- 1 file changed, 27 insertions(+), 4 deletions(-) diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 06f04c4c3d..5d443dbfcb 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -358,6 +358,28 @@ let unify ?loc env t1 t2 = user_err ?loc (str "This expression has type " ++ pr_glbtype name t1 ++ str " but an expression what expected of type " ++ pr_glbtype name t2) +let unify_arrow ?loc env ft args = + let ft0 = ft in + let rec iter ft args is_fun = match kind env ft, args with + | t, [] -> t + | GTypArrow (t1, ft), (loc, t2) :: args -> + let () = unify ~loc env t1 t2 in + iter ft args true + | GTypVar id, (_, t) :: args -> + let ft = GTypVar (fresh_id env) in + let () = unify_var env id (GTypArrow (t, ft)) in + iter ft args true + | (GTypRef _ | GTypTuple _), _ :: _ -> + let name = env_name env in + if is_fun then + user_err ?loc (str "This function has type " ++ pr_glbtype name ft0 ++ + str " and is applied to too many arguments") + else + user_err ?loc (str "This expression has type " ++ pr_glbtype name ft0 ++ + str " and is not a function") + in + iter ft args false + (** Term typing *) let is_pure_constructor kn = @@ -637,14 +659,15 @@ let rec intern_rec env = function let loc = loc_of_tacexpr e in intern_constructor env loc kn args | CTacApp (loc, f, args) -> + let loc = loc_of_tacexpr f in let (f, ft) = intern_rec env f in let fold arg (args, t) = + let loc = loc_of_tacexpr arg in let (arg, argt) = intern_rec env arg in - (arg :: args, GTypArrow (argt, t)) + (arg :: args, (loc, argt) :: t) in - let ret = GTypVar (fresh_id env) in - let (args, t) = List.fold_right fold args ([], ret) in - let () = unify ~loc env ft t in + let (args, t) = List.fold_right fold args ([], []) in + let ret = unify_arrow ~loc env ft t in (GTacApp (f, args), ret) | CTacLet (loc, false, el, e) -> let fold accu ((loc, na), _, _) = match na with -- cgit v1.2.3 From cfb181899cdd076fb7f2e061089ba76067e47ccc Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 26 Jul 2017 16:34:09 +0200 Subject: Fix typo in error message --- src/tac2intern.ml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 5d443dbfcb..79e33f3a94 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -356,14 +356,14 @@ let unify ?loc env t1 t2 = with CannotUnify (u1, u2) -> let name = env_name env in user_err ?loc (str "This expression has type " ++ pr_glbtype name t1 ++ - str " but an expression what expected of type " ++ pr_glbtype name t2) + str " but an expression was expected of type " ++ pr_glbtype name t2) let unify_arrow ?loc env ft args = let ft0 = ft in let rec iter ft args is_fun = match kind env ft, args with | t, [] -> t | GTypArrow (t1, ft), (loc, t2) :: args -> - let () = unify ~loc env t1 t2 in + let () = unify ~loc env t2 t1 in iter ft args true | GTypVar id, (_, t) :: args -> let ft = GTypVar (fresh_id env) in -- cgit v1.2.3 From 2a74da7b6f275634fd8ed9c209edc73f2ae15427 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 26 Jul 2017 16:38:52 +0200 Subject: Adding a file for testing typing. --- Makefile | 4 ++ tests/Makefile | 7 +++ tests/ltac2.v | 131 ---------------------------------------------------- tests/stuff/ltac2.v | 131 ++++++++++++++++++++++++++++++++++++++++++++++++++++ tests/typing.v | 25 ++++++++++ 5 files changed, 167 insertions(+), 131 deletions(-) create mode 100644 tests/Makefile delete mode 100644 tests/ltac2.v create mode 100644 tests/stuff/ltac2.v create mode 100644 tests/typing.v diff --git a/Makefile b/Makefile index cfdeeba747..d555fea236 100644 --- a/Makefile +++ b/Makefile @@ -10,3 +10,7 @@ clean: Makefile.coq Makefile.coq: _CoqProject $(COQBIN)/coq_makefile -f _CoqProject -o Makefile.coq + +tests: all + @$(MAKE) -C tests -s clean + @$(MAKE) -C tests -s all diff --git a/tests/Makefile b/tests/Makefile new file mode 100644 index 0000000000..a48ab0860f --- /dev/null +++ b/tests/Makefile @@ -0,0 +1,7 @@ +all: $(patsubst %.v,%.v.log,$(wildcard *.v)) + +%.v.log: %.v + $(COQBIN)/coqtop -I ../src -Q ../theories Ltac2 < $< 2> $@ + +clean: + rm -f *.log diff --git a/tests/ltac2.v b/tests/ltac2.v deleted file mode 100644 index 770d385406..0000000000 --- a/tests/ltac2.v +++ /dev/null @@ -1,131 +0,0 @@ -Require Import Ltac2.Ltac2. - -Ltac2 foo (_ : int) := - let f (x : int) := x in - let _ := f 0 in - f 1. - -Print Ltac2 foo. - -Import Control. - -Ltac2 exact x := refine (fun _ => x). - -Print Ltac2 refine. -Print Ltac2 exact. - -Ltac2 foo' _ := ident:(bla). - -Print Ltac2 foo'. - -Ltac2 bar x H := match x with -| None => constr:(fun H => ltac2:(exact (hyp ident:(H))) -> nat) -| Some x => x -end. - -Print Ltac2 bar. - -Ltac2 qux := Some 0. - -Print Ltac2 qux. - -Ltac2 Type foo := [ Foo (int) ]. - -Fail Ltac2 qux0 := Foo None. - -Ltac2 Type 'a ref := { mutable contents : 'a }. - -Fail Ltac2 qux0 := { contents := None }. -Ltac2 foo0 _ := { contents := None }. - -Print Ltac2 foo0. - -Ltac2 qux0 x := x.(contents). -Ltac2 qux1 x := x.(contents) := x.(contents). - -Ltac2 qux2 := ([1;2], true). - -Print Ltac2 qux0. -Print Ltac2 qux1. -Print Ltac2 qux2. - -Import Control. - -Ltac2 qux3 x := constr:(nat -> ltac2:(refine (fun _ => hyp x))). - -Print Ltac2 qux3. - -Ltac2 qux4 f x := x, (f x, x). - -Print Ltac2 qux4. - -Ltac2 Type rec nat := [ O | S (nat) ]. - -Ltac2 message_of_nat n := -let rec aux n := -match n with -| O => Message.of_string "O" -| S n => Message.concat (Message.of_string "S") (aux n) -end in aux n. - -Print Ltac2 message_of_nat. - -Ltac2 numgoals (_ : unit) := - let r := { contents := O } in - enter (fun _ => r.(contents) := S (r.(contents))); - r.(contents). - -Print Ltac2 numgoals. - -Goal True /\ False. -Proof. -let n := numgoals () in Message.print (message_of_nat n). -refine (fun _ => open_constr:((fun x => conj _ _) 0)); (). -let n := numgoals () in Message.print (message_of_nat n). - -Fail (hyp ident:(x)). -Fail (enter (fun _ => hyp ident:(There_is_no_spoon); ())). - -enter (fun _ => Message.print (Message.of_string "foo")). - -enter (fun _ => Message.print (Message.of_constr (goal ()))). -Fail enter (fun _ => Message.print (Message.of_constr (qux3 ident:(x)))). -enter (fun _ => plus (fun _ => constr:(_); ()) (fun _ => ())). -plus - (fun _ => enter (fun _ => let x := ident:(foo) in let _ := hyp x in ())) (fun _ => Message.print (Message.of_string "failed")). -let x := { contents := 0 } in -let x := x.(contents) := x.(contents) in x. -Abort. - -Ltac2 Type exn ::= [ Foo ]. - -Goal True. -Proof. -plus (fun _ => zero Foo) (fun _ => ()). -Abort. - -Ltac2 Type exn ::= [ Bar (string) ]. - -Goal True. -Proof. -Fail zero (Bar "lol"). -Abort. - -Ltac2 Notation "refine!" c(constr) := refine c. - -Goal True. -Proof. -refine! I. -Abort. - -Goal True. -Proof. -let x _ := plus (fun _ => 0) (fun _ => 1) in -match case x with -| Val x => - match x with - | (x, k) => Message.print (Message.of_int (k Not_found)) - end -| Err x => Message.print (Message.of_string "Err") -end. -Abort. diff --git a/tests/stuff/ltac2.v b/tests/stuff/ltac2.v new file mode 100644 index 0000000000..770d385406 --- /dev/null +++ b/tests/stuff/ltac2.v @@ -0,0 +1,131 @@ +Require Import Ltac2.Ltac2. + +Ltac2 foo (_ : int) := + let f (x : int) := x in + let _ := f 0 in + f 1. + +Print Ltac2 foo. + +Import Control. + +Ltac2 exact x := refine (fun _ => x). + +Print Ltac2 refine. +Print Ltac2 exact. + +Ltac2 foo' _ := ident:(bla). + +Print Ltac2 foo'. + +Ltac2 bar x H := match x with +| None => constr:(fun H => ltac2:(exact (hyp ident:(H))) -> nat) +| Some x => x +end. + +Print Ltac2 bar. + +Ltac2 qux := Some 0. + +Print Ltac2 qux. + +Ltac2 Type foo := [ Foo (int) ]. + +Fail Ltac2 qux0 := Foo None. + +Ltac2 Type 'a ref := { mutable contents : 'a }. + +Fail Ltac2 qux0 := { contents := None }. +Ltac2 foo0 _ := { contents := None }. + +Print Ltac2 foo0. + +Ltac2 qux0 x := x.(contents). +Ltac2 qux1 x := x.(contents) := x.(contents). + +Ltac2 qux2 := ([1;2], true). + +Print Ltac2 qux0. +Print Ltac2 qux1. +Print Ltac2 qux2. + +Import Control. + +Ltac2 qux3 x := constr:(nat -> ltac2:(refine (fun _ => hyp x))). + +Print Ltac2 qux3. + +Ltac2 qux4 f x := x, (f x, x). + +Print Ltac2 qux4. + +Ltac2 Type rec nat := [ O | S (nat) ]. + +Ltac2 message_of_nat n := +let rec aux n := +match n with +| O => Message.of_string "O" +| S n => Message.concat (Message.of_string "S") (aux n) +end in aux n. + +Print Ltac2 message_of_nat. + +Ltac2 numgoals (_ : unit) := + let r := { contents := O } in + enter (fun _ => r.(contents) := S (r.(contents))); + r.(contents). + +Print Ltac2 numgoals. + +Goal True /\ False. +Proof. +let n := numgoals () in Message.print (message_of_nat n). +refine (fun _ => open_constr:((fun x => conj _ _) 0)); (). +let n := numgoals () in Message.print (message_of_nat n). + +Fail (hyp ident:(x)). +Fail (enter (fun _ => hyp ident:(There_is_no_spoon); ())). + +enter (fun _ => Message.print (Message.of_string "foo")). + +enter (fun _ => Message.print (Message.of_constr (goal ()))). +Fail enter (fun _ => Message.print (Message.of_constr (qux3 ident:(x)))). +enter (fun _ => plus (fun _ => constr:(_); ()) (fun _ => ())). +plus + (fun _ => enter (fun _ => let x := ident:(foo) in let _ := hyp x in ())) (fun _ => Message.print (Message.of_string "failed")). +let x := { contents := 0 } in +let x := x.(contents) := x.(contents) in x. +Abort. + +Ltac2 Type exn ::= [ Foo ]. + +Goal True. +Proof. +plus (fun _ => zero Foo) (fun _ => ()). +Abort. + +Ltac2 Type exn ::= [ Bar (string) ]. + +Goal True. +Proof. +Fail zero (Bar "lol"). +Abort. + +Ltac2 Notation "refine!" c(constr) := refine c. + +Goal True. +Proof. +refine! I. +Abort. + +Goal True. +Proof. +let x _ := plus (fun _ => 0) (fun _ => 1) in +match case x with +| Val x => + match x with + | (x, k) => Message.print (Message.of_int (k Not_found)) + end +| Err x => Message.print (Message.of_string "Err") +end. +Abort. diff --git a/tests/typing.v b/tests/typing.v new file mode 100644 index 0000000000..8460ab42b7 --- /dev/null +++ b/tests/typing.v @@ -0,0 +1,25 @@ +Require Import Ltac2.Ltac2. + +(** Ltac2 is typed à la ML. *) + +Ltac2 test0 n := Int.add n 1. + +Print Ltac2 test0. + +Ltac2 test1 () := test0 0. + +Print Ltac2 test1. + +Fail Ltac2 test2 () := test0 true. + +Fail Ltac2 test2 () := test0 0 0. + +(** Polymorphism *) + +Ltac2 rec list_length l := +match l with +| [] => 0 +| x :: l => Int.add 1 (list_length l) +end. + +Print Ltac2 list_length. -- cgit v1.2.3 From 4395637a6471fc95934fe93da671bda68d415a77 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 26 Jul 2017 17:53:04 +0200 Subject: Ensuring that inductive constructors are always capitalized. --- src/g_ltac2.ml4 | 23 +++++++--- src/tac2core.ml | 6 +-- src/tac2entries.ml | 29 +++++++------ src/tac2env.ml | 75 ++++++++++++++++++-------------- src/tac2env.mli | 17 ++++++-- src/tac2expr.mli | 9 ++-- src/tac2intern.ml | 120 +++++++++++++++++++--------------------------------- src/tac2print.ml | 4 +- tests/stuff/ltac2.v | 2 +- 9 files changed, 144 insertions(+), 141 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 605cb75d66..88a64dacd9 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -28,20 +28,32 @@ let inj_constr loc c = inj_wit Stdarg.wit_constr loc c let inj_open_constr loc c = inj_wit Stdarg.wit_open_constr loc c let inj_ident loc c = inj_wit Stdarg.wit_ident loc c +let pattern_of_qualid loc id = + if Tac2env.is_constructor (snd id) then CPatRef (loc, RelId id, []) + else + let (dp, id) = Libnames.repr_qualid (snd id) in + if DirPath.is_empty dp then CPatVar (Some loc, Name id) + else + CErrors.user_err ~loc (Pp.str "Syntax error") + GEXTEND Gram GLOBAL: tac2expr tac2type tac2def_val tac2def_typ tac2def_ext tac2def_syn; tac2pat: [ "1" LEFTA - [ id = Prim.qualid; pl = LIST1 tac2pat LEVEL "0" -> CPatRef (!@loc, RelId id, pl) - | id = Prim.qualid -> CPatRef (!@loc, RelId id, []) + [ id = Prim.qualid; pl = LIST1 tac2pat LEVEL "0" -> + if Tac2env.is_constructor (snd id) then + CPatRef (!@loc, RelId id, pl) + else + CErrors.user_err ~loc:!@loc (Pp.str "Syntax error") + | id = Prim.qualid -> pattern_of_qualid !@loc id | "["; "]" -> CPatRef (!@loc, AbsKn Tac2core.Core.c_nil, []) | p1 = tac2pat; "::"; p2 = tac2pat -> CPatRef (!@loc, AbsKn Tac2core.Core.c_cons, [p1; p2]) ] | "0" - [ "_" -> CPatAny (!@loc) + [ "_" -> CPatVar (Some !@loc, Anonymous) | "()" -> CPatTup (Loc.tag ~loc:!@loc []) - | id = Prim.qualid -> CPatRef (!@loc, RelId id, []) + | id = Prim.qualid -> pattern_of_qualid !@loc id | "("; pl = LIST0 tac2pat LEVEL "1" SEP ","; ")" -> CPatTup (Loc.tag ~loc:!@loc pl) ] ] ; @@ -90,7 +102,8 @@ GEXTEND Gram tactic_atom: [ [ n = Prim.integer -> CTacAtm (Loc.tag ~loc:!@loc (AtmInt n)) | s = Prim.string -> CTacAtm (Loc.tag ~loc:!@loc (AtmStr s)) - | id = Prim.qualid -> CTacRef (RelId id) + | id = Prim.qualid -> + if Tac2env.is_constructor (snd id) then CTacCst (RelId id) else CTacRef (RelId id) | "@"; id = Prim.ident -> inj_ident !@loc id | "'"; c = Constr.constr -> inj_open_constr !@loc c | IDENT "constr"; ":"; "("; c = Constr.lconstr; ")" -> inj_constr !@loc c diff --git a/src/tac2core.ml b/src/tac2core.ml index b665f761ce..2ccc49b043 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -734,7 +734,7 @@ let dummy_loc = Loc.make_loc (-1, -1) let rthunk e = let loc = Tac2intern.loc_of_tacexpr e in - let var = [CPatAny loc, Some (CTypRef (loc, AbsKn Core.t_unit, []))] in + let var = [CPatVar (Some loc, Anonymous), Some (CTypRef (loc, AbsKn Core.t_unit, []))] in CTacFun (loc, var, e) let add_generic_scope s entry arg = @@ -795,9 +795,9 @@ let () = add_scope "opt" begin function let scope = Extend.Aopt scope in let act opt = match opt with | None -> - CTacRef (AbsKn (TacConstructor Core.c_none)) + CTacCst (AbsKn Core.c_none) | Some x -> - CTacApp (dummy_loc, CTacRef (AbsKn (TacConstructor Core.c_some)), [act x]) + CTacApp (dummy_loc, CTacCst (AbsKn Core.c_some), [act x]) in Tac2entries.ScopeRule (scope, act) | _ -> scope_fail () diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 100041f15e..da0e213340 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -34,14 +34,14 @@ type tacdef = { } let perform_tacdef visibility ((sp, kn), def) = - let () = if not def.tacdef_local then Tac2env.push_ltac visibility sp (TacConstant kn) in + let () = if not def.tacdef_local then Tac2env.push_ltac visibility sp kn in Tac2env.define_global kn (def.tacdef_expr, def.tacdef_type) let load_tacdef i obj = perform_tacdef (Until i) obj let open_tacdef i obj = perform_tacdef (Exactly i) obj let cache_tacdef ((sp, kn), def) = - let () = Tac2env.push_ltac (Until 1) sp (TacConstant kn) in + let () = Tac2env.push_ltac (Until 1) sp kn in Tac2env.define_global kn (def.tacdef_expr, def.tacdef_type) let subst_tacdef (subst, def) = @@ -83,7 +83,7 @@ let push_typedef visibility sp kn (_, def) = match def with let iter (c, _) = let spc = change_sp_label sp c in let knc = change_kn_label kn c in - Tac2env.push_ltac visibility spc (TacConstructor knc) + Tac2env.push_constructor visibility spc knc in Tac2env.push_type visibility sp kn; List.iter iter cstrs @@ -185,7 +185,7 @@ let push_typext vis sp kn def = let iter data = let spc = change_sp_label sp data.edata_name in let knc = change_kn_label kn data.edata_name in - Tac2env.push_ltac vis spc (TacConstructor knc) + Tac2env.push_constructor vis spc knc in List.iter iter def.typext_expr @@ -620,12 +620,18 @@ let register_struct ?local str = match str with let print_ltac ref = let (loc, qid) = qualid_of_reference ref in - let kn = - try Tac2env.locate_ltac qid - with Not_found -> user_err ?loc (str "Unknown tactic " ++ pr_qualid qid) - in - match kn with - | TacConstant kn -> + if Tac2env.is_constructor qid then + let kn = + try Tac2env.locate_constructor qid + with Not_found -> user_err ?loc (str "Unknown constructor " ++ pr_qualid qid) + in + let _ = Tac2env.interp_constructor kn in + Feedback.msg_notice (hov 2 (str "Constructor" ++ spc () ++ str ":" ++ spc () ++ pr_qualid qid)) + else + let kn = + try Tac2env.locate_ltac qid + with Not_found -> user_err ?loc (str "Unknown tactic " ++ pr_qualid qid) + in let (e, _, (_, t)) = Tac2env.interp_global kn in let name = int_name () in Feedback.msg_notice ( @@ -634,9 +640,6 @@ let print_ltac ref = hov 2 (pr_qualid qid ++ spc () ++ str ":=" ++ spc () ++ pr_glbexpr e) ) ) - | TacConstructor kn -> - let _ = Tac2env.interp_constructor kn in - Feedback.msg_notice (hov 2 (str "Constructor" ++ spc () ++ str ":" ++ spc () ++ pr_qualid qid)) (** Calling tactics *) diff --git a/src/tac2env.ml b/src/tac2env.ml index 08c7b321be..6e47e78a9d 100644 --- a/src/tac2env.ml +++ b/src/tac2env.ml @@ -115,25 +115,13 @@ struct id, (DirPath.repr dir) end -type tacref = Tac2expr.tacref = -| TacConstant of ltac_constant -| TacConstructor of ltac_constructor - -module TacRef = -struct -type t = tacref -let equal r1 r2 = match r1, r2 with -| TacConstant c1, TacConstant c2 -> KerName.equal c1 c2 -| TacConstructor c1, TacConstructor c2 -> KerName.equal c1 c2 -| _ -> false -end - module KnTab = Nametab.Make(FullPath)(KerName) -module RfTab = Nametab.Make(FullPath)(TacRef) type nametab = { - tab_ltac : RfTab.t; - tab_ltac_rev : full_path KNmap.t * full_path KNmap.t; + tab_ltac : KnTab.t; + tab_ltac_rev : full_path KNmap.t; + tab_cstr : KnTab.t; + tab_cstr_rev : full_path KNmap.t; tab_type : KnTab.t; tab_type_rev : full_path KNmap.t; tab_proj : KnTab.t; @@ -141,8 +129,10 @@ type nametab = { } let empty_nametab = { - tab_ltac = RfTab.empty; - tab_ltac_rev = (KNmap.empty, KNmap.empty); + tab_ltac = KnTab.empty; + tab_ltac_rev = KNmap.empty; + tab_cstr = KnTab.empty; + tab_cstr_rev = KNmap.empty; tab_type = KnTab.empty; tab_type_rev = KNmap.empty; tab_proj = KnTab.empty; @@ -153,29 +143,41 @@ let nametab = Summary.ref empty_nametab ~name:"ltac2-nametab" let push_ltac vis sp kn = let tab = !nametab in - let tab_ltac = RfTab.push vis sp kn tab.tab_ltac in - let (constant_map, constructor_map) = tab.tab_ltac_rev in - let tab_ltac_rev = match kn with - | TacConstant c -> (KNmap.add c sp constant_map, constructor_map) - | TacConstructor c -> (constant_map, KNmap.add c sp constructor_map) - in + let tab_ltac = KnTab.push vis sp kn tab.tab_ltac in + let tab_ltac_rev = KNmap.add kn sp tab.tab_ltac_rev in nametab := { tab with tab_ltac; tab_ltac_rev } let locate_ltac qid = let tab = !nametab in - RfTab.locate qid tab.tab_ltac + KnTab.locate qid tab.tab_ltac let locate_extended_all_ltac qid = let tab = !nametab in - RfTab.find_prefixes qid tab.tab_ltac + KnTab.find_prefixes qid tab.tab_ltac let shortest_qualid_of_ltac kn = let tab = !nametab in - let sp = match kn with - | TacConstant c -> KNmap.find c (fst tab.tab_ltac_rev) - | TacConstructor c -> KNmap.find c (snd tab.tab_ltac_rev) - in - RfTab.shortest_qualid Id.Set.empty sp tab.tab_ltac + let sp = KNmap.find kn tab.tab_ltac_rev in + KnTab.shortest_qualid Id.Set.empty sp tab.tab_ltac + +let push_constructor vis sp kn = + let tab = !nametab in + let tab_cstr = KnTab.push vis sp kn tab.tab_cstr in + let tab_cstr_rev = KNmap.add kn sp tab.tab_cstr_rev in + nametab := { tab with tab_cstr; tab_cstr_rev } + +let locate_constructor qid = + let tab = !nametab in + KnTab.locate qid tab.tab_cstr + +let locate_extended_all_constructor qid = + let tab = !nametab in + KnTab.find_prefixes qid tab.tab_cstr + +let shortest_qualid_of_constructor kn = + let tab = !nametab in + let sp = KNmap.find kn tab.tab_cstr_rev in + KnTab.shortest_qualid Id.Set.empty sp tab.tab_cstr let push_type vis sp kn = let tab = !nametab in @@ -240,3 +242,14 @@ let coq_prefix = (** Generic arguments *) let wit_ltac2 = Genarg.make0 "ltac2" + +let is_constructor qid = + let (_, id) = repr_qualid qid in + let id = Id.to_string id in + assert (String.length id > 0); + match id with + | "true" | "false" -> true (** built-in constructors *) + | _ -> + match id.[0] with + | 'A'..'Z' -> true + | _ -> false diff --git a/src/tac2env.mli b/src/tac2env.mli index c4b8c1e0ca..8ab9656cb9 100644 --- a/src/tac2env.mli +++ b/src/tac2env.mli @@ -63,10 +63,15 @@ val interp_projection : ltac_projection -> projection_data (** {5 Name management} *) -val push_ltac : visibility -> full_path -> tacref -> unit -val locate_ltac : qualid -> tacref -val locate_extended_all_ltac : qualid -> tacref list -val shortest_qualid_of_ltac : tacref -> qualid +val push_ltac : visibility -> full_path -> ltac_constant -> unit +val locate_ltac : qualid -> ltac_constant +val locate_extended_all_ltac : qualid -> ltac_constant list +val shortest_qualid_of_ltac : ltac_constant -> qualid + +val push_constructor : visibility -> full_path -> ltac_constructor -> unit +val locate_constructor : qualid -> ltac_constructor +val locate_extended_all_constructor : qualid -> ltac_constructor list +val shortest_qualid_of_constructor : ltac_constructor -> qualid val push_type : visibility -> full_path -> type_constant -> unit val locate_type : qualid -> type_constant @@ -104,3 +109,7 @@ val coq_prefix : ModPath.t (** {5 Generic arguments} *) val wit_ltac2 : (raw_tacexpr, glb_tacexpr, Util.Empty.t) genarg_type + +(** {5 Helper functions} *) + +val is_constructor : qualid -> bool diff --git a/src/tac2expr.mli b/src/tac2expr.mli index a9f2109cb2..b268e70cb3 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -70,19 +70,16 @@ type atom = | AtmInt of int | AtmStr of string -type tacref = -| TacConstant of ltac_constant -| TacConstructor of ltac_constructor - (** Tactic expressions *) type raw_patexpr = -| CPatAny of Loc.t +| CPatVar of Name.t located | CPatRef of Loc.t * ltac_constructor or_relid * raw_patexpr list | CPatTup of raw_patexpr list located type raw_tacexpr = | CTacAtm of atom located -| CTacRef of tacref or_relid +| CTacRef of ltac_constant or_relid +| CTacCst of ltac_constructor or_relid | CTacFun of Loc.t * (raw_patexpr * raw_typexpr option) list * raw_tacexpr | CTacApp of Loc.t * raw_tacexpr * raw_tacexpr list | CTacLet of Loc.t * rec_flag * (Name.t located * raw_typexpr option * raw_tacexpr) list * raw_tacexpr diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 79e33f3a94..3ea35171bb 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -191,6 +191,8 @@ let loc_of_tacexpr = function | CTacAtm (loc, _) -> Option.default dummy_loc loc | CTacRef (RelId (loc, _)) -> Option.default dummy_loc loc | CTacRef (AbsKn _) -> dummy_loc +| CTacCst (RelId (loc, _)) -> Option.default dummy_loc loc +| CTacCst (AbsKn _) -> dummy_loc | CTacFun (loc, _, _) -> loc | CTacApp (loc, _, _) -> loc | CTacLet (loc, _, _, _) -> loc @@ -206,7 +208,7 @@ let loc_of_tacexpr = function | CTacExt (loc, _) -> loc let loc_of_patexpr = function -| CPatAny loc -> loc +| CPatVar (loc, _) -> Option.default dummy_loc loc | CPatRef (loc, _, _) -> loc | CPatTup (loc, _) -> Option.default dummy_loc loc @@ -516,22 +518,17 @@ let get_variable env var = let get_constructor env var = match var with | RelId (loc, qid) -> - let c = try Some (Tac2env.locate_ltac qid) with Not_found -> None in + let c = try Some (Tac2env.locate_constructor qid) with Not_found -> None in begin match c with - | Some (TacConstructor knc) -> + | Some knc -> let kn = Tac2env.interp_constructor knc in - ArgArg (kn, knc) - | Some (TacConstant _) -> - CErrors.user_err ?loc (str "The term " ++ pr_qualid qid ++ - str " is not the constructor of an inductive type.") + (kn, knc) | None -> - let (dp, id) = repr_qualid qid in - if DirPath.is_empty dp then ArgVar (loc, id) - else CErrors.user_err ?loc (str "Unbound constructor " ++ pr_qualid qid) + CErrors.user_err ?loc (str "Unbound constructor " ++ pr_qualid qid) end | AbsKn knc -> let kn = Tac2env.interp_constructor knc in - ArgArg (kn, knc) + (kn, knc) let get_projection var = match var with | RelId (loc, qid) -> @@ -562,18 +559,10 @@ type glb_patexpr = | GPatTup of glb_patexpr list let rec intern_patexpr env = function -| CPatAny _ -> GPatVar Anonymous -| CPatRef (_, qid, []) -> - begin match get_constructor env qid with - | ArgVar (_, id) -> GPatVar (Name id) - | ArgArg (_, kn) -> GPatRef (kn, []) - end +| CPatVar (_, na) -> GPatVar na | CPatRef (_, qid, pl) -> - begin match get_constructor env qid with - | ArgVar (loc, id) -> - user_err ?loc (str "Unbound constructor " ++ Nameops.pr_id id) - | ArgArg (_, kn) -> GPatRef (kn, List.map (fun p -> intern_patexpr env p) pl) - end + let (_, kn) = get_constructor env qid in + GPatRef (kn, List.map (fun p -> intern_patexpr env p) pl) | CPatTup (_, pl) -> GPatTup (List.map (fun p -> intern_patexpr env p) pl) @@ -603,10 +592,6 @@ let get_pattern_kind env pl = match pl with (** Internalization *) -let is_constructor env qid = match get_variable env qid with -| ArgArg (TacConstructor _) -> true -| _ -> false - (** Used to generate a fresh tactic variable for pattern-expansion *) let fresh_var env = let bad id = @@ -617,18 +602,19 @@ let fresh_var env = let rec intern_rec env = function | CTacAtm (_, atm) -> intern_atm env atm -| CTacRef qid as e -> +| CTacRef qid -> begin match get_variable env qid with | ArgVar (_, id) -> let sch = Id.Map.find id env.env_var in (GTacVar id, fresh_mix_type_scheme env sch) - | ArgArg (TacConstant kn) -> + | ArgArg kn -> let (_, _, sch) = Tac2env.interp_global kn in (GTacRef kn, fresh_type_scheme env sch) - | ArgArg (TacConstructor kn) -> - let loc = loc_of_tacexpr e in - intern_constructor env loc kn [] end +| CTacCst qid as e -> + let loc = loc_of_tacexpr e in + let (_, kn) = get_constructor env qid in + intern_constructor env loc kn [] | CTacFun (loc, bnd, e) -> let fold (env, bnd, tl) (pat, t) = let t = match t with @@ -651,11 +637,8 @@ let rec intern_rec env = function let (e, t) = intern_rec env e in let t = List.fold_left (fun accu t -> GTypArrow (t, accu)) t tl in (GTacFun (bnd, e), t) -| CTacApp (loc, CTacRef qid, args) as e when is_constructor env qid -> - let kn = match get_variable env qid with - | ArgArg (TacConstructor kn) -> kn - | _ -> assert false - in +| CTacApp (loc, CTacCst qid, args) as e -> + let (_, kn) = get_constructor env qid in let loc = loc_of_tacexpr e in intern_constructor env loc kn args | CTacApp (loc, f, args) -> @@ -823,7 +806,7 @@ and intern_let_rec env loc el e = to depth-one where leaves are either variables or catch-all *) and intern_case env loc e pl = let (e', t) = intern_rec env e in - let todo ~loc () = user_err ~loc (str "Pattern not handled yet") in + let todo ?loc () = user_err ?loc (str "Pattern not handled yet") in match get_pattern_kind env pl with | PKind_any -> let (pat, b) = List.hd pl in @@ -848,12 +831,7 @@ and intern_case env loc e pl = (GTacCse (e', GCaseAlg t_unit, [|b|], [||]), tb) | [CPatTup (_, pl), b] -> let map = function - | CPatAny _ -> Anonymous - | CPatRef (loc, qid, []) -> - begin match get_constructor env qid with - | ArgVar (_, id) -> Name id - | ArgArg _ -> todo ~loc () - end + | CPatVar (_, na) -> na | p -> todo ~loc:(loc_of_patexpr p) () in let ids = List.map map pl in @@ -885,7 +863,8 @@ and intern_case env loc e pl = | [] -> () | (pat, br) :: rem -> let tbr = match pat with - | CPatAny _ -> + | CPatVar (loc, Name _) -> todo ?loc () + | CPatVar (_, Anonymous) -> let () = check_redundant_clause rem in let (br', brT) = intern_rec env br in (** Fill all remaining branches *) @@ -906,23 +885,14 @@ and intern_case env loc e pl = let _ = List.fold_left fill (0, 0) cstrs in brT | CPatRef (loc, qid, args) -> - let data = match get_constructor env qid with - | ArgVar _ -> todo ~loc () - | ArgArg (data, _) -> - let () = - let kn' = data.cdata_type in - if not (KerName.equal kn kn') then - invalid_pattern ~loc kn (GCaseAlg kn') - in - data + let (data, _) = get_constructor env qid in + let () = + let kn' = data.cdata_type in + if not (KerName.equal kn kn') then + invalid_pattern ~loc kn (GCaseAlg kn') in let get_id = function - | CPatAny _ -> Anonymous - | CPatRef (loc, qid, []) -> - begin match get_constructor env qid with - | ArgVar (_, id) -> Name id - | ArgArg _ -> todo ~loc () - end + | CPatVar (_, na) -> na | p -> todo ~loc:(loc_of_patexpr p) () in let ids = List.map get_id args in @@ -1165,12 +1135,9 @@ let get_projection0 var = match var with | AbsKn kn -> kn let rec ids_of_pattern accu = function -| CPatAny _ -> accu -| CPatRef (_, RelId (_, qid), pl) -> - let (dp, id) = repr_qualid qid in - let accu = if DirPath.is_empty dp then Id.Set.add id accu else accu in - List.fold_left ids_of_pattern accu pl -| CPatRef (_, AbsKn _, pl) -> +| CPatVar (_, Anonymous) -> accu +| CPatVar (_, Name id) -> Id.Set.add id accu +| CPatRef (_, _, pl) -> List.fold_left ids_of_pattern accu pl | CPatTup (_, pl) -> List.fold_left ids_of_pattern accu pl @@ -1183,6 +1150,9 @@ let rec globalize ids e = match e with | ArgVar _ -> e | ArgArg kn -> CTacRef (AbsKn kn) end +| CTacCst qid -> + let (_, knc) = get_constructor () qid in + CTacCst (AbsKn knc) | CTacFun (loc, bnd, e) -> let fold (pats, accu) (pat, t) = let accu = ids_of_pattern accu pat in @@ -1252,12 +1222,10 @@ and globalize_case ids (p, e) = (globalize_pattern ids p, globalize ids e) and globalize_pattern ids p = match p with -| CPatAny _ -> p +| CPatVar _ -> p | CPatRef (loc, cst, pl) -> - let cst = match get_constructor () cst with - | ArgVar _ -> cst - | ArgArg (_, knc) -> AbsKn knc - in + let (_, knc) = get_constructor () cst in + let cst = AbsKn knc in let pl = List.map (fun p -> globalize_pattern ids p) pl in CPatRef (loc, cst, pl) | CPatTup (loc, pl) -> @@ -1393,12 +1361,9 @@ let rec subst_rawtype subst t = match t with let subst_tacref subst ref = match ref with | RelId _ -> ref -| AbsKn (TacConstant kn) -> - let kn' = subst_kn subst kn in - if kn' == kn then ref else AbsKn (TacConstant kn') -| AbsKn (TacConstructor kn) -> +| AbsKn kn -> let kn' = subst_kn subst kn in - if kn' == kn then ref else AbsKn (TacConstructor kn') + if kn' == kn then ref else AbsKn kn' let subst_projection subst prj = match prj with | RelId _ -> prj @@ -1407,7 +1372,7 @@ let subst_projection subst prj = match prj with if kn' == kn then prj else AbsKn kn' let rec subst_rawpattern subst p = match p with -| CPatAny _ -> p +| CPatVar _ -> p | CPatRef (loc, c, pl) -> let pl' = List.smartmap (fun p -> subst_rawpattern subst p) pl in let c' = match c with @@ -1427,6 +1392,9 @@ let rec subst_rawexpr subst t = match t with | CTacRef ref -> let ref' = subst_tacref subst ref in if ref' == ref then t else CTacRef ref' +| CTacCst ref -> + let ref' = subst_tacref subst ref in + if ref' == ref then t else CTacCst ref' | CTacFun (loc, bnd, e) -> let map (na, t as p) = let t' = Option.smartmap (fun t -> subst_rawtype subst t) t in diff --git a/src/tac2print.ml b/src/tac2print.ml index e6f0582e3d..2afcfb4a6e 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -83,7 +83,7 @@ let int_name () = (** Term printing *) let pr_constructor kn = - Libnames.pr_qualid (Tac2env.shortest_qualid_of_ltac (TacConstructor kn)) + Libnames.pr_qualid (Tac2env.shortest_qualid_of_constructor kn) let pr_projection kn = Libnames.pr_qualid (Tac2env.shortest_qualid_of_projection kn) @@ -138,7 +138,7 @@ let pr_glbexpr_gen lvl c = | GTacAtm atm -> pr_atom atm | GTacVar id -> Id.print id | GTacRef gr -> - let qid = shortest_qualid_of_ltac (TacConstant gr) in + let qid = shortest_qualid_of_ltac gr in Libnames.pr_qualid qid | GTacFun (nas, c) -> let nas = pr_sequence pr_name nas in diff --git a/tests/stuff/ltac2.v b/tests/stuff/ltac2.v index 770d385406..36ea1ef4c1 100644 --- a/tests/stuff/ltac2.v +++ b/tests/stuff/ltac2.v @@ -18,7 +18,7 @@ Ltac2 foo' _ := ident:(bla). Print Ltac2 foo'. -Ltac2 bar x H := match x with +Ltac2 bar x h := match x with | None => constr:(fun H => ltac2:(exact (hyp ident:(H))) -> nat) | Some x => x end. -- cgit v1.2.3 From e917841e46264ad7b80241b25dcd7731eca468a8 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 26 Jul 2017 19:05:22 +0200 Subject: Do not expand trivial patterns in functions. --- src/tac2intern.ml | 27 ++++++++++++++++++--------- 1 file changed, 18 insertions(+), 9 deletions(-) diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 3ea35171bb..431b551191 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -616,27 +616,36 @@ let rec intern_rec env = function let (_, kn) = get_constructor env qid in intern_constructor env loc kn [] | CTacFun (loc, bnd, e) -> + (** Expand pattern: [fun p => t] becomes [fun x => match x with p => t end] *) let fold (env, bnd, tl) (pat, t) = let t = match t with | None -> GTypVar (fresh_id env) | Some t -> intern_type env t in - let id = fresh_var env in - let env = push_name (Name id) (monomorphic t) env in - (env, (id, pat) :: bnd, t :: tl) + let na, expand = match pat with + | CPatVar (_, na) -> + (** Don't expand variable patterns *) + na, None + | _ -> + let loc = loc_of_patexpr pat in + let id = fresh_var env in + let qid = RelId (Loc.tag ~loc (qualid_of_ident id)) in + Name id, Some qid + in + let env = push_name na (monomorphic t) env in + (env, (na, pat, expand) :: bnd, t :: tl) in let (env, bnd, tl) = List.fold_left fold (env, [], []) bnd in - (** Expand pattern: [fun p => t] becomes [fun x => match x with p => t end] *) - let fold e (id, pat) = - let loc = loc_of_patexpr pat in - let qid = RelId (Loc.tag ~loc (qualid_of_ident id)) in + let fold e (na, pat, expand) = match expand with + | None -> e + | Some qid -> CTacCse (loc, CTacRef qid, [pat, e]) in let e = List.fold_left fold e bnd in - let bnd = List.rev_map (fun (id, _) -> Name id) bnd in + let nas = List.rev_map (fun (na, _, _) -> na) bnd in let (e, t) = intern_rec env e in let t = List.fold_left (fun accu t -> GTypArrow (t, accu)) t tl in - (GTacFun (bnd, e), t) + (GTacFun (nas, e), t) | CTacApp (loc, CTacCst qid, args) as e -> let (_, kn) = get_constructor env qid in let loc = loc_of_tacexpr e in -- cgit v1.2.3 From 57b9df4e07351a753f897dc24eb8238f6465b26d Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 26 Jul 2017 20:50:58 +0200 Subject: Dedicated module for ident type. --- _CoqProject | 1 + src/tac2core.ml | 29 +++++++++++++++++++++++++++++ theories/Ident.v | 17 +++++++++++++++++ theories/Ltac2.v | 1 + 4 files changed, 48 insertions(+) create mode 100644 theories/Ident.v diff --git a/_CoqProject b/_CoqProject index 6d3470cfa7..2561b7b6ec 100644 --- a/_CoqProject +++ b/_CoqProject @@ -23,6 +23,7 @@ src/ltac2_plugin.mlpack theories/Init.v theories/Int.v theories/String.v +theories/Ident.v theories/Array.v theories/Control.v theories/Message.v diff --git a/src/tac2core.ml b/src/tac2core.ml index 2ccc49b043..f29cc8c89e 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -267,6 +267,31 @@ let prm_array_get : ml_tactic = function else wrap (fun () -> v.(n)) | _ -> assert false +(** Ident *) + +let prm_ident_equal : ml_tactic = function +| [id1; id2] -> + let id1 = Value.to_ident id1 in + let id2 = Value.to_ident id2 in + return (Value.of_bool (Id.equal id1 id2)) +| _ -> assert false + +let prm_ident_to_string : ml_tactic = function +| [id] -> + let id = Value.to_ident id in + return (Value.of_string (Id.to_string id)) +| _ -> assert false + +let prm_ident_of_string : ml_tactic = function +| [s] -> + let s = Value.to_string s in + let id = + try Value.of_option (Some (Value.of_ident (Id.of_string s))) + with _ -> Value.of_option None + in + return id +| _ -> assert false + (** Int *) let prm_int_equal : ml_tactic = function @@ -624,6 +649,10 @@ let () = Tac2env.define_primitive (pname "int_add") prm_int_add let () = Tac2env.define_primitive (pname "int_sub") prm_int_sub let () = Tac2env.define_primitive (pname "int_mul") prm_int_mul +let () = Tac2env.define_primitive (pname "ident_equal") prm_ident_equal +let () = Tac2env.define_primitive (pname "ident_to_string") prm_ident_to_string +let () = Tac2env.define_primitive (pname "ident_of_string") prm_ident_of_string + let () = Tac2env.define_primitive (pname "throw") prm_throw let () = Tac2env.define_primitive (pname "zero") prm_zero diff --git a/theories/Ident.v b/theories/Ident.v new file mode 100644 index 0000000000..55456afbe2 --- /dev/null +++ b/theories/Ident.v @@ -0,0 +1,17 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* t -> bool := "ltac2" "ident_equal". + +Ltac2 @ external of_string : string -> t option := "ltac2" "ident_of_string". + +Ltac2 @ external to_string : t -> string := "ltac2" "ident_to_string". diff --git a/theories/Ltac2.v b/theories/Ltac2.v index 4cd3fafcb2..0d6c8f232a 100644 --- a/theories/Ltac2.v +++ b/theories/Ltac2.v @@ -10,6 +10,7 @@ Require Export Ltac2.Init. Require Ltac2.Int. Require Ltac2.String. +Require Ltac2.Ident. Require Ltac2.Array. Require Ltac2.Message. Require Ltac2.Constr. -- cgit v1.2.3 From 2d6461140fadf1af8b92567b77e869eb2bd9dd7e Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 26 Jul 2017 22:41:55 +0200 Subject: Tentative fix of parsing of product types. --- src/g_ltac2.ml4 | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 88a64dacd9..cd227a78f1 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -121,7 +121,7 @@ GEXTEND Gram [ "5" RIGHTA [ t1 = tac2type; "->"; t2 = tac2type -> CTypArrow (!@loc, t1, t2) ] | "2" - [ t = tac2type; "*"; tl = LIST1 tac2type SEP "*" -> CTypTuple (!@loc, t :: tl) ] + [ t = tac2type; "*"; tl = LIST1 tac2type LEVEL "1" SEP "*" -> CTypTuple (!@loc, t :: tl) ] | "1" LEFTA [ t = SELF; qid = Prim.qualid -> CTypRef (!@loc, RelId qid, [t]) ] | "0" -- cgit v1.2.3 From c9e7d7f1ceb22667e77a4ee49a4afc2cce6f9a2c Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 26 Jul 2017 22:47:18 +0200 Subject: Adding an example file --- tests/example1.v | 26 ++++++++++++++++++++++++++ 1 file changed, 26 insertions(+) create mode 100644 tests/example1.v diff --git a/tests/example1.v b/tests/example1.v new file mode 100644 index 0000000000..1b26aad824 --- /dev/null +++ b/tests/example1.v @@ -0,0 +1,26 @@ +Require Import Ltac2.Ltac2. + +Import Ltac2.Control. + +(** Alternative implementation of the hyp primitive *) +Ltac2 get_hyp_by_name x := + let h := hyps () in + let rec find x l := match l with + | [] => zero Not_found + | p :: l => + match p with + | (id, _, t) => + match Ident.equal x id with + | true => t + | false => find x l + end + end + end in + find x h. + +Print Ltac2 get_hyp_by_name. + +Goal forall n m, n + m = 0 -> n = 0. +Proof. +refine (fun () => '(fun n m H => _)). +let t := get_hyp_by_name @H in Message.print (Message.of_constr t). -- cgit v1.2.3 From f204058d329fa78b506f4c3b3c4f97ecce504d4b Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 26 Jul 2017 21:27:36 +0200 Subject: Adding necessary primitives to do pattern-matching over constr. --- _CoqProject | 1 + src/g_ltac2.ml4 | 2 ++ src/tac2core.ml | 81 ++++++++++++++++++++++++++++++++++++++++++++++++++++++ src/tac2env.ml | 1 + src/tac2env.mli | 2 ++ theories/Init.v | 4 +++ theories/Ltac2.v | 1 + theories/Pattern.v | 30 ++++++++++++++++++++ 8 files changed, 122 insertions(+) create mode 100644 theories/Pattern.v diff --git a/_CoqProject b/_CoqProject index 2561b7b6ec..e3ad3987bd 100644 --- a/_CoqProject +++ b/_CoqProject @@ -28,5 +28,6 @@ theories/Array.v theories/Control.v theories/Message.v theories/Constr.v +theories/Pattern.v theories/Std.v theories/Ltac2.v diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index cd227a78f1..d7d376f88a 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -27,6 +27,7 @@ let inj_wit wit loc x = CTacExt (loc, Genarg.in_gen (Genarg.rawwit wit) x) let inj_constr loc c = inj_wit Stdarg.wit_constr loc c let inj_open_constr loc c = inj_wit Stdarg.wit_open_constr loc c let inj_ident loc c = inj_wit Stdarg.wit_ident loc c +let inj_pattern loc c = inj_wit Tac2env.wit_pattern loc c let pattern_of_qualid loc id = if Tac2env.is_constructor (snd id) then CPatRef (loc, RelId id, []) @@ -109,6 +110,7 @@ GEXTEND Gram | IDENT "constr"; ":"; "("; c = Constr.lconstr; ")" -> inj_constr !@loc c | IDENT "open_constr"; ":"; "("; c = Constr.lconstr; ")" -> inj_open_constr !@loc c | IDENT "ident"; ":"; "("; c = Prim.ident; ")" -> inj_ident !@loc c + | IDENT "pattern"; ":"; "("; c = Constr.lconstr_pattern; ")" -> inj_pattern !@loc c ] ] ; let_clause: diff --git a/src/tac2core.ml b/src/tac2core.ml index f29cc8c89e..ab1eaec9d9 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -26,6 +26,7 @@ let val_tag t = match val_tag t with let val_constr = val_tag (topwit Stdarg.wit_constr) let val_ident = val_tag (topwit Stdarg.wit_ident) +let val_pattern = Val.create "ltac2:pattern" let val_pp = Val.create "ltac2:pp" let val_sort = Val.create "ltac2:sort" let val_cast = Val.create "ltac2:cast" @@ -51,6 +52,7 @@ let t_array = coq_core "array" let t_unit = coq_core "unit" let t_list = coq_core "list" let t_constr = coq_core "constr" +let t_pattern = coq_core "pattern" let t_ident = coq_core "ident" let t_option = coq_core "option" @@ -121,6 +123,9 @@ let to_constr c = to_ext val_constr c let of_ident c = of_ext val_ident c let to_ident c = to_ext val_ident c +let of_pattern c = of_ext val_pattern c +let to_pattern c = to_ext val_pattern c + (** FIXME: handle backtrace in Ltac2 exceptions *) let of_exn c = match fst c with | LtacError (kn, c) -> ValOpn (kn, c) @@ -178,6 +183,9 @@ let err_outofbounds = let err_notfound = LtacError (coq_core "Not_found", [||]) +let err_matchfailure = + LtacError (coq_core "Match_failure", [||]) + (** Helper functions *) let thaw f = interp_app f [v_unit] @@ -464,6 +472,54 @@ let prm_constr_kind : ml_tactic = function end | _ -> assert false +(** Patterns *) + +let prm_pattern_matches : ml_tactic = function +| [pat; c] -> + let pat = Value.to_pattern pat in + let c = Value.to_constr c in + pf_apply begin fun env sigma -> + let ans = + try Some (Constr_matching.matches env sigma pat c) + with Constr_matching.PatternMatchingFailure -> None + in + begin match ans with + | None -> Proofview.tclZERO err_matchfailure + | Some ans -> + let ans = Id.Map.bindings ans in + let of_pair (id, c) = Value.of_tuple [| Value.of_ident id; Value.of_constr c |] in + return (Value.of_list (List.map of_pair ans)) + end + end +| _ -> assert false + +let prm_pattern_matches_subterm : ml_tactic = function +| [pat; c] -> + let pat = Value.to_pattern pat in + let c = Value.to_constr c in + let open Constr_matching in + let rec of_ans s = match IStream.peek s with + | IStream.Nil -> Proofview.tclZERO err_matchfailure + | IStream.Cons ({ m_sub = (_, sub); m_ctx }, s) -> + let ans = Id.Map.bindings sub in + let of_pair (id, c) = Value.of_tuple [| Value.of_ident id; Value.of_constr c |] in + let ans = Value.of_tuple [| Value.of_constr m_ctx; Value.of_list (List.map of_pair ans) |] in + Proofview.tclOR (return ans) (fun _ -> of_ans s) + in + pf_apply begin fun env sigma -> + let ans = Constr_matching.match_appsubterm env sigma pat c in + of_ans ans + end +| _ -> assert false + +let prm_pattern_instantiate : ml_tactic = function +| [ctx; c] -> + let ctx = EConstr.Unsafe.to_constr (Value.to_constr ctx) in + let c = EConstr.Unsafe.to_constr (Value.to_constr c) in + let ans = Termops.subst_meta [Constr_matching.special_meta, c] ctx in + return (Value.of_constr (EConstr.of_constr ans)) +| _ -> assert false + (** Error *) let prm_throw : ml_tactic = function @@ -642,6 +698,10 @@ let () = Tac2env.define_primitive (pname "constr_type") prm_constr_type let () = Tac2env.define_primitive (pname "constr_equal") prm_constr_equal let () = Tac2env.define_primitive (pname "constr_kind") prm_constr_kind +let () = Tac2env.define_primitive (pname "pattern_matches") prm_pattern_matches +let () = Tac2env.define_primitive (pname "pattern_matches_subterm") prm_pattern_matches_subterm +let () = Tac2env.define_primitive (pname "pattern_instantiate") prm_pattern_instantiate + let () = Tac2env.define_primitive (pname "int_equal") prm_int_equal let () = Tac2env.define_primitive (pname "int_compare") prm_int_compare let () = Tac2env.define_primitive (pname "int_neg") prm_int_neg @@ -738,6 +798,14 @@ let () = } in define_ml_object Stdarg.wit_ident obj +let () = + let interp _ c = return (Val.Dyn (val_pattern, c)) in + let obj = { + ml_type = t_pattern; + ml_interp = interp; + } in + define_ml_object Tac2env.wit_pattern obj + let () = let interp ist env sigma concl tac = let fold id (Val.Dyn (tag, v)) (accu : environment) : environment = @@ -752,6 +820,19 @@ let () = in Pretyping.register_constr_interp0 wit_ltac2 interp +(** Patterns *) + +let () = + let intern ist c = + let _, pat = Constrintern.intern_constr_pattern ist.Genintern.genv ~as_type:false c in + ist, pat + in + Genintern.register_intern0 wit_pattern intern + +let () = + let subst s c = Patternops.subst_pattern s c in + Genintern.register_subst0 wit_pattern subst + (** Built-in notation scopes *) let add_scope s f = diff --git a/src/tac2env.ml b/src/tac2env.ml index 6e47e78a9d..2094898ced 100644 --- a/src/tac2env.ml +++ b/src/tac2env.ml @@ -242,6 +242,7 @@ let coq_prefix = (** Generic arguments *) let wit_ltac2 = Genarg.make0 "ltac2" +let wit_pattern = Genarg.make0 "ltac2:pattern" let is_constructor qid = let (_, id) = repr_qualid qid in diff --git a/src/tac2env.mli b/src/tac2env.mli index 8ab9656cb9..e26109b691 100644 --- a/src/tac2env.mli +++ b/src/tac2env.mli @@ -110,6 +110,8 @@ val coq_prefix : ModPath.t val wit_ltac2 : (raw_tacexpr, glb_tacexpr, Util.Empty.t) genarg_type +val wit_pattern : (Constrexpr.constr_expr, Pattern.constr_pattern, Pattern.constr_pattern) genarg_type + (** {5 Helper functions} *) val is_constructor : qualid -> bool diff --git a/theories/Init.v b/theories/Init.v index c0a73576d3..803e48e352 100644 --- a/theories/Init.v +++ b/theories/Init.v @@ -27,6 +27,7 @@ Ltac2 Type constant. Ltac2 Type inductive. Ltac2 Type constructor. Ltac2 Type projection. +Ltac2 Type pattern. Ltac2 Type constr. Ltac2 Type message. @@ -56,3 +57,6 @@ Ltac2 Type exn ::= [ Not_focussed ]. Ltac2 Type exn ::= [ Not_found ]. (** Used when something is missing. *) + +Ltac2 Type exn ::= [ Match_failure ]. +(** Used to signal a pattern didn't match a term. *) diff --git a/theories/Ltac2.v b/theories/Ltac2.v index 0d6c8f232a..9aaee850cd 100644 --- a/theories/Ltac2.v +++ b/theories/Ltac2.v @@ -15,4 +15,5 @@ Require Ltac2.Array. Require Ltac2.Message. Require Ltac2.Constr. Require Ltac2.Control. +Require Ltac2.Pattern. Require Ltac2.Std. diff --git a/theories/Pattern.v b/theories/Pattern.v new file mode 100644 index 0000000000..c2ba4162e8 --- /dev/null +++ b/theories/Pattern.v @@ -0,0 +1,30 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* constr -> (ident * constr) list := + "ltac2" "pattern_matches". +(** If the term matches the pattern, returns the bound variables. If it doesn't, + fail with [Match_failure]. Panics if not focussed. *) + +Ltac2 @ external matches_subterm : t -> constr -> context * ((ident * constr) list) := + "ltac2" "pattern_matches_subterm". +(** Returns a stream of results corresponding to all of the subterms of the term + that matches the pattern as in [matches]. The stream is encoded as a + backtracking value whose last exception is [Match_failure]. The additional + value compared to [matches] is the context of the match, to be filled with + the instantiate function. *) + +Ltac2 @ external instantiate : context -> constr -> constr := + "ltac2" "pattern_instantiate". +(** Fill the hole of a context with the given term. *) -- cgit v1.2.3 From 6b3fd33d7e3b775ce6afe38f7b16d4b11bdccdb3 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 27 Jul 2017 16:05:34 +0200 Subject: Fix expansion of toplevel let-rec after the constructor / constant split. --- src/tac2entries.ml | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/src/tac2entries.ml b/src/tac2entries.ml index da0e213340..6fce739d30 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -269,8 +269,7 @@ let inline_rec_tactic tactics = let map_body ((loc, id), _, e) = (loc, Name id), None, e in let bnd = List.map map_body tactics in let pat_of_id (loc, id) = - let qid = (loc, qualid_of_ident id) in - (CPatRef (Option.default dummy_loc loc, RelId qid, []), None) + (CPatVar (loc, Name id), None) in let var_of_id (loc, id) = let qid = (loc, qualid_of_ident id) in -- cgit v1.2.3 From bc9a18bd48fb43a2aedd9c11df7a3e4244074120 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 27 Jul 2017 16:20:16 +0200 Subject: Using thunks in the horrible Ltac2 example. --- tests/stuff/ltac2.v | 37 ++++++++++++++++++++++--------------- 1 file changed, 22 insertions(+), 15 deletions(-) diff --git a/tests/stuff/ltac2.v b/tests/stuff/ltac2.v index 36ea1ef4c1..86ab5afb17 100644 --- a/tests/stuff/ltac2.v +++ b/tests/stuff/ltac2.v @@ -9,12 +9,12 @@ Print Ltac2 foo. Import Control. -Ltac2 exact x := refine (fun _ => x). +Ltac2 exact x := refine (fun () => x). Print Ltac2 refine. Print Ltac2 exact. -Ltac2 foo' _ := ident:(bla). +Ltac2 foo' () := ident:(bla). Print Ltac2 foo'. @@ -36,7 +36,7 @@ Fail Ltac2 qux0 := Foo None. Ltac2 Type 'a ref := { mutable contents : 'a }. Fail Ltac2 qux0 := { contents := None }. -Ltac2 foo0 _ := { contents := None }. +Ltac2 foo0 () := { contents := None }. Print Ltac2 foo0. @@ -51,7 +51,7 @@ Print Ltac2 qux2. Import Control. -Ltac2 qux3 x := constr:(nat -> ltac2:(refine (fun _ => hyp x))). +Ltac2 qux3 x := constr:(nat -> ltac2:(refine (fun () => hyp x))). Print Ltac2 qux3. @@ -70,9 +70,9 @@ end in aux n. Print Ltac2 message_of_nat. -Ltac2 numgoals (_ : unit) := +Ltac2 numgoals () := let r := { contents := O } in - enter (fun _ => r.(contents) := S (r.(contents))); + enter (fun () => r.(contents) := S (r.(contents))); r.(contents). Print Ltac2 numgoals. @@ -80,19 +80,19 @@ Print Ltac2 numgoals. Goal True /\ False. Proof. let n := numgoals () in Message.print (message_of_nat n). -refine (fun _ => open_constr:((fun x => conj _ _) 0)); (). +refine (fun () => open_constr:((fun x => conj _ _) 0)); (). let n := numgoals () in Message.print (message_of_nat n). Fail (hyp ident:(x)). -Fail (enter (fun _ => hyp ident:(There_is_no_spoon); ())). +Fail (enter (fun () => hyp ident:(There_is_no_spoon); ())). -enter (fun _ => Message.print (Message.of_string "foo")). +enter (fun () => Message.print (Message.of_string "foo")). -enter (fun _ => Message.print (Message.of_constr (goal ()))). -Fail enter (fun _ => Message.print (Message.of_constr (qux3 ident:(x)))). -enter (fun _ => plus (fun _ => constr:(_); ()) (fun _ => ())). +enter (fun () => Message.print (Message.of_constr (goal ()))). +Fail enter (fun () => Message.print (Message.of_constr (qux3 ident:(x)))). +enter (fun () => plus (fun () => constr:(_); ()) (fun _ => ())). plus - (fun _ => enter (fun _ => let x := ident:(foo) in let _ := hyp x in ())) (fun _ => Message.print (Message.of_string "failed")). + (fun () => enter (fun () => let x := ident:(foo) in let _ := hyp x in ())) (fun _ => Message.print (Message.of_string "failed")). let x := { contents := 0 } in let x := x.(contents) := x.(contents) in x. Abort. @@ -101,7 +101,7 @@ Ltac2 Type exn ::= [ Foo ]. Goal True. Proof. -plus (fun _ => zero Foo) (fun _ => ()). +plus (fun () => zero Foo) (fun _ => ()). Abort. Ltac2 Type exn ::= [ Bar (string) ]. @@ -120,7 +120,7 @@ Abort. Goal True. Proof. -let x _ := plus (fun _ => 0) (fun _ => 1) in +let x () := plus (fun () => 0) (fun _ => 1) in match case x with | Val x => match x with @@ -129,3 +129,10 @@ match case x with | Err x => Message.print (Message.of_string "Err") end. Abort. + +Ltac2 rec do n tac := match Int.equal n 0 with +| true => () +| false => tac (); do (Int.sub n 1) tac +end. + +Print Ltac2 do. -- cgit v1.2.3 From fee254e2f7a343629df31d5a39780ca4698de571 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 27 Jul 2017 16:29:27 +0200 Subject: Cleaning up code in internalization. --- src/tac2entries.ml | 16 +++++++++++++--- src/tac2expr.mli | 11 ++++++++++- src/tac2intern.ml | 33 ++++++++++++++++++++------------- src/tac2print.ml | 4 ++-- 4 files changed, 45 insertions(+), 19 deletions(-) diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 6fce739d30..5490f9d57f 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -78,7 +78,7 @@ let change_sp_label sp id = let push_typedef visibility sp kn (_, def) = match def with | GTydDef _ -> Tac2env.push_type visibility sp kn -| GTydAlg cstrs -> +| GTydAlg { galg_constructors = cstrs } -> (** Register constructors *) let iter (c, _) = let spc = change_sp_label sp c in @@ -107,7 +107,7 @@ let next i = let define_typedef kn (params, def as qdef) = match def with | GTydDef _ -> Tac2env.define_type kn qdef -| GTydAlg cstrs -> +| GTydAlg { galg_constructors = cstrs } -> (** Define constructors *) let constant = ref 0 in let nonconstant = ref 0 in @@ -665,7 +665,17 @@ let call ~default e = let register_prim_alg name params def = let id = Id.of_string name in let def = List.map (fun (cstr, tpe) -> (Id.of_string_soft cstr, tpe)) def in - let def = (params, GTydAlg def) in + let getn (const, nonconst) (c, args) = match args with + | [] -> (succ const, nonconst) + | _ :: _ -> (const, succ nonconst) + in + let nconst, nnonconst = List.fold_left getn (0, 0) def in + let alg = { + galg_constructors = def; + galg_nconst = nconst; + galg_nnonconst = nnonconst; + } in + let def = (params, GTydAlg alg) in let def = { typdef_local = false; typdef_expr = def } in ignore (Lib.add_leaf id (inTypDef def)) diff --git a/src/tac2expr.mli b/src/tac2expr.mli index b268e70cb3..e564dbde78 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -53,9 +53,18 @@ type 'a glb_typexpr = | GTypTuple of 'a glb_typexpr list | GTypRef of type_constant * 'a glb_typexpr list +type glb_alg_type = { + galg_constructors : (uid * int glb_typexpr list) list; + (** Constructors of the algebraic type *) + galg_nconst : int; + (** Number of constant constructors *) + galg_nnonconst : int; + (** Number of non-constant constructors *) +} + type glb_typedef = | GTydDef of int glb_typexpr option -| GTydAlg of (uid * int glb_typexpr list) list +| GTydAlg of glb_alg_type | GTydRec of (lid * mutable_flag * int glb_typexpr) list | GTydOpn diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 431b551191..86db803ea7 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -482,7 +482,7 @@ let check_elt_empty loc env t = match kind env t with | GTypRef (kn, _) -> let def = Tac2env.interp_type kn in match def with - | _, GTydAlg [] -> kn + | _, GTydAlg { galg_constructors = [] } -> kn | _ -> let name = env_name env in user_err ~loc (str "Type " ++ pr_glbtype name t ++ str " is not an empty type") @@ -856,17 +856,13 @@ and intern_case env loc e pl = let subst, tc = fresh_reftype env kn in let () = unify ~loc:(loc_of_tacexpr e) env t tc in let (params, def) = Tac2env.interp_type kn in - let cstrs = match def with + let galg = match def with | GTydAlg c -> c | _ -> assert false in - let count (const, nonconst) (c, args) = match args with - | [] -> (succ const, nonconst) - | _ :: _ -> (const, succ nonconst) - in - let nconst, nnonconst = List.fold_left count (0, 0) cstrs in - let const = Array.make nconst None in - let nonconst = Array.make nnonconst None in + let cstrs = galg.galg_constructors in + let const = Array.make galg.galg_nconst None in + let nonconst = Array.make galg.galg_nnonconst None in let ret = GTypVar (fresh_id env) in let rec intern_branch = function | [] -> () @@ -1114,7 +1110,17 @@ let intern_typedef self (ids, t) : glb_quant_typedef = | CTydAlg constrs -> let map (c, t) = (c, List.map intern t) in let constrs = List.map map constrs in - (count, GTydAlg constrs) + let getn (const, nonconst) (c, args) = match args with + | [] -> (succ const, nonconst) + | _ :: _ -> (const, succ nonconst) + in + let nconst, nnonconst = List.fold_left getn (0, 0) constrs in + let galg = { + galg_constructors = constrs; + galg_nconst = nconst; + galg_nnonconst = nnonconst; + } in + (count, GTydAlg galg) | CTydRec fields -> let map (c, mut, t) = (c, mut, intern t) in let fields = List.map map fields in @@ -1324,13 +1330,14 @@ let subst_typedef subst e = match e with | GTydDef t -> let t' = Option.smartmap (fun t -> subst_type subst t) t in if t' == t then e else GTydDef t' -| GTydAlg constrs -> +| GTydAlg galg -> let map (c, tl as p) = let tl' = List.smartmap (fun t -> subst_type subst t) tl in if tl' == tl then p else (c, tl') in - let constrs' = List.smartmap map constrs in - if constrs' == constrs then e else GTydAlg constrs' + let constrs' = List.smartmap map galg.galg_constructors in + if constrs' == galg.galg_constructors then e + else GTydAlg { galg with galg_constructors = constrs' } | GTydRec fields -> let map (c, mut, t as p) = let t' = subst_type subst t in diff --git a/src/tac2print.ml b/src/tac2print.ml index 2afcfb4a6e..aa8e1e98d5 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -175,7 +175,7 @@ let pr_glbexpr_gen lvl c = mt () (** FIXME when implemented *) | GTacCst (GCaseAlg tpe, n, cl) -> begin match Tac2env.interp_type tpe with - | _, GTydAlg def -> + | _, GTydAlg { galg_constructors = def } -> let paren = match lvl with | E0 -> paren | E1 | E2 | E3 | E4 | E5 -> fun x -> x @@ -202,7 +202,7 @@ let pr_glbexpr_gen lvl c = let br = match info with | GCaseAlg kn -> let def = match Tac2env.interp_type kn with - | _, GTydAlg def -> def + | _, GTydAlg { galg_constructors = def } -> def | _, GTydDef _ | _, GTydRec _ | _, GTydOpn -> assert false in let br = order_branches cst_br ncst_br def in -- cgit v1.2.3 From 86e7ec3bd7b26b1d377c8397b62346f5e44f5d87 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 27 Jul 2017 16:50:25 +0200 Subject: Factorizing code for constructors and tuples. --- src/g_ltac2.ml4 | 23 ++-- src/tac2core.ml | 6 +- src/tac2entries.ml | 11 +- src/tac2expr.mli | 20 ++-- src/tac2intern.ml | 321 ++++++++++++++++++++++------------------------------ src/tac2intern.mli | 2 +- src/tac2print.ml | 24 ++-- tests/stuff/ltac2.v | 1 + 8 files changed, 186 insertions(+), 222 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index d7d376f88a..7ee9d7e282 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -47,15 +47,16 @@ GEXTEND Gram else CErrors.user_err ~loc:!@loc (Pp.str "Syntax error") | id = Prim.qualid -> pattern_of_qualid !@loc id - | "["; "]" -> CPatRef (!@loc, AbsKn Tac2core.Core.c_nil, []) + | "["; "]" -> CPatRef (!@loc, AbsKn (Other Tac2core.Core.c_nil), []) | p1 = tac2pat; "::"; p2 = tac2pat -> - CPatRef (!@loc, AbsKn Tac2core.Core.c_cons, [p1; p2]) + CPatRef (!@loc, AbsKn (Other Tac2core.Core.c_cons), [p1; p2]) ] | "0" [ "_" -> CPatVar (Some !@loc, Anonymous) - | "()" -> CPatTup (Loc.tag ~loc:!@loc []) + | "()" -> CPatRef (!@loc, AbsKn (Tuple 0), []) | id = Prim.qualid -> pattern_of_qualid !@loc id - | "("; pl = LIST0 tac2pat LEVEL "1" SEP ","; ")" -> CPatTup (Loc.tag ~loc:!@loc pl) ] + | "("; pl = LIST0 tac2pat LEVEL "1" SEP ","; ")" -> + CPatRef (!@loc, AbsKn (Tuple (List.length pl)), pl) ] ] ; tac2expr: @@ -73,12 +74,14 @@ GEXTEND Gram [ e = tac2expr; el = LIST1 tac2expr LEVEL "0" -> CTacApp (!@loc, e, el) | e = SELF; ".("; qid = Prim.qualid; ")" -> CTacPrj (!@loc, e, RelId qid) | e = SELF; ".("; qid = Prim.qualid; ")"; ":="; r = tac2expr LEVEL "1" -> CTacSet (!@loc, e, RelId qid, r) - | e0 = tac2expr; ","; el = LIST1 tac2expr LEVEL "1" SEP "," -> CTacTup (Loc.tag ~loc:!@loc (e0 :: el)) ] + | e0 = tac2expr; ","; el = LIST1 tac2expr LEVEL "1" SEP "," -> + let el = e0 :: el in + CTacApp (!@loc, CTacCst (!@loc, AbsKn (Tuple (List.length el))), el) ] | "0" [ "("; a = tac2expr LEVEL "5"; ")" -> a | "("; a = tac2expr; ":"; t = tac2type; ")" -> CTacCnv (!@loc, a, t) - | "()" -> CTacTup (Loc.tag ~loc:!@loc []) - | "("; ")" -> CTacTup (Loc.tag ~loc:!@loc []) + | "()" -> CTacCst (!@loc, AbsKn (Tuple 0)) + | "("; ")" -> CTacCst (!@loc, AbsKn (Tuple 0)) | "["; a = LIST0 tac2expr LEVEL "1" SEP ";"; "]" -> CTacLst (Loc.tag ~loc:!@loc a) | "{"; a = tac2rec_fieldexprs; "}" -> CTacRec (!@loc, a) | a = tactic_atom -> a ] @@ -104,7 +107,7 @@ GEXTEND Gram [ [ n = Prim.integer -> CTacAtm (Loc.tag ~loc:!@loc (AtmInt n)) | s = Prim.string -> CTacAtm (Loc.tag ~loc:!@loc (AtmStr s)) | id = Prim.qualid -> - if Tac2env.is_constructor (snd id) then CTacCst (RelId id) else CTacRef (RelId id) + if Tac2env.is_constructor (snd id) then CTacCst (!@loc, RelId id) else CTacRef (RelId id) | "@"; id = Prim.ident -> inj_ident !@loc id | "'"; c = Constr.constr -> inj_open_constr !@loc c | IDENT "constr"; ":"; "("; c = Constr.lconstr; ")" -> inj_constr !@loc c @@ -123,7 +126,9 @@ GEXTEND Gram [ "5" RIGHTA [ t1 = tac2type; "->"; t2 = tac2type -> CTypArrow (!@loc, t1, t2) ] | "2" - [ t = tac2type; "*"; tl = LIST1 tac2type LEVEL "1" SEP "*" -> CTypTuple (!@loc, t :: tl) ] + [ t = tac2type; "*"; tl = LIST1 tac2type LEVEL "1" SEP "*" -> + let tl = t :: tl in + CTypRef (!@loc, AbsKn (Tuple (List.length tl)), tl) ] | "1" LEFTA [ t = SELF; qid = Prim.qualid -> CTypRef (!@loc, RelId qid, [t]) ] | "0" diff --git a/src/tac2core.ml b/src/tac2core.ml index ab1eaec9d9..a3678d1ad0 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -844,7 +844,7 @@ let dummy_loc = Loc.make_loc (-1, -1) let rthunk e = let loc = Tac2intern.loc_of_tacexpr e in - let var = [CPatVar (Some loc, Anonymous), Some (CTypRef (loc, AbsKn Core.t_unit, []))] in + let var = [CPatVar (Some loc, Anonymous), Some (CTypRef (loc, AbsKn (Other Core.t_unit), []))] in CTacFun (loc, var, e) let add_generic_scope s entry arg = @@ -905,9 +905,9 @@ let () = add_scope "opt" begin function let scope = Extend.Aopt scope in let act opt = match opt with | None -> - CTacCst (AbsKn Core.c_none) + CTacCst (dummy_loc, AbsKn (Other Core.c_none)) | Some x -> - CTacApp (dummy_loc, CTacCst (AbsKn Core.c_some), [act x]) + CTacApp (dummy_loc, CTacCst (dummy_loc, AbsKn (Other Core.c_some)), [act x]) in Tac2entries.ScopeRule (scope, act) | _ -> scope_fail () diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 5490f9d57f..70f1b583e0 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -655,7 +655,7 @@ let solve default tac = let call ~default e = let loc = loc_of_tacexpr e in - let (e, (_, t)) = intern e in + let (e, t) = intern e in let () = check_unit ~loc t in let tac = Tac2interp.interp Id.Map.empty e in solve default (Proofview.tclIGNORE tac) @@ -681,12 +681,17 @@ let register_prim_alg name params def = let coq_def n = KerName.make2 Tac2env.coq_prefix (Label.make n) +let def_unit = { + typdef_local = false; + typdef_expr = 0, GTydDef (Some (GTypRef (Tuple 0, []))); +} + let t_list = coq_def "list" let _ = Mltop.declare_cache_obj begin fun () -> - register_prim_alg "unit" 0 ["()", []]; + ignore (Lib.add_leaf (Id.of_string "unit") (inTypDef def_unit)); register_prim_alg "list" 1 [ ("[]", []); - ("::", [GTypVar 0; GTypRef (t_list, [GTypVar 0])]); + ("::", [GTypVar 0; GTypRef (Other t_list, [GTypVar 0])]); ]; end "ltac2_plugin" diff --git a/src/tac2expr.mli b/src/tac2expr.mli index e564dbde78..668bc0dad1 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -33,13 +33,16 @@ type ml_tactic_name = { mltac_tactic : string; } +type 'a or_tuple = +| Tuple of int +| Other of 'a + (** {5 Type syntax} *) type raw_typexpr = | CTypVar of Name.t located | CTypArrow of Loc.t * raw_typexpr * raw_typexpr -| CTypTuple of Loc.t * raw_typexpr list -| CTypRef of Loc.t * type_constant or_relid * raw_typexpr list +| CTypRef of Loc.t * type_constant or_tuple or_relid * raw_typexpr list type raw_typedef = | CTydDef of raw_typexpr option @@ -50,8 +53,7 @@ type raw_typedef = type 'a glb_typexpr = | GTypVar of 'a | GTypArrow of 'a glb_typexpr * 'a glb_typexpr -| GTypTuple of 'a glb_typexpr list -| GTypRef of type_constant * 'a glb_typexpr list +| GTypRef of type_constant or_tuple * 'a glb_typexpr list type glb_alg_type = { galg_constructors : (uid * int glb_typexpr list) list; @@ -82,17 +84,15 @@ type atom = (** Tactic expressions *) type raw_patexpr = | CPatVar of Name.t located -| CPatRef of Loc.t * ltac_constructor or_relid * raw_patexpr list -| CPatTup of raw_patexpr list located +| CPatRef of Loc.t * ltac_constructor or_tuple or_relid * raw_patexpr list type raw_tacexpr = | CTacAtm of atom located | CTacRef of ltac_constant or_relid -| CTacCst of ltac_constructor or_relid +| CTacCst of Loc.t * ltac_constructor or_tuple or_relid | CTacFun of Loc.t * (raw_patexpr * raw_typexpr option) list * raw_tacexpr | CTacApp of Loc.t * raw_tacexpr * raw_tacexpr list | CTacLet of Loc.t * rec_flag * (Name.t located * raw_typexpr option * raw_tacexpr) list * raw_tacexpr -| CTacTup of raw_tacexpr list located | CTacArr of raw_tacexpr list located | CTacLst of raw_tacexpr list located | CTacCnv of Loc.t * raw_tacexpr * raw_typexpr @@ -107,9 +107,7 @@ and raw_taccase = raw_patexpr * raw_tacexpr and raw_recexpr = (ltac_projection or_relid * raw_tacexpr) list -type case_info = -| GCaseTuple of int -| GCaseAlg of type_constant +type case_info = type_constant or_tuple type 'a open_match = { opn_match : 'a; diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 86db803ea7..e460111fc1 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -27,8 +27,8 @@ let t_array = coq_type "array" let t_unit = coq_type "unit" let t_list = coq_type "list" -let c_nil = GTacCst (GCaseAlg t_list, 0, []) -let c_cons e el = GTacCst (GCaseAlg t_list, 0, [e; el]) +let c_nil = GTacCst (Other t_list, 0, []) +let c_cons e el = GTacCst (Other t_list, 0, [e; el]) (** Union find *) @@ -191,12 +191,10 @@ let loc_of_tacexpr = function | CTacAtm (loc, _) -> Option.default dummy_loc loc | CTacRef (RelId (loc, _)) -> Option.default dummy_loc loc | CTacRef (AbsKn _) -> dummy_loc -| CTacCst (RelId (loc, _)) -> Option.default dummy_loc loc -| CTacCst (AbsKn _) -> dummy_loc +| CTacCst (loc, _) -> loc | CTacFun (loc, _, _) -> loc | CTacApp (loc, _, _) -> loc | CTacLet (loc, _, _, _) -> loc -| CTacTup (loc, _) -> Option.default dummy_loc loc | CTacArr (loc, _) -> Option.default dummy_loc loc | CTacLst (loc, _) -> Option.default dummy_loc loc | CTacCnv (loc, _, _) -> loc @@ -210,7 +208,6 @@ let loc_of_tacexpr = function let loc_of_patexpr = function | CPatVar (loc, _) -> Option.default dummy_loc loc | CPatRef (loc, _, _) -> loc -| CPatTup (loc, _) -> Option.default dummy_loc loc let error_nargs_mismatch loc nargs nfound = user_err ~loc (str "Constructor expects " ++ int nargs ++ @@ -225,7 +222,6 @@ let error_nparams_mismatch loc nargs nfound = let rec subst_type subst (t : 'a glb_typexpr) = match t with | GTypVar id -> subst id | GTypArrow (t1, t2) -> GTypArrow (subst_type subst t1, subst_type subst t2) -| GTypTuple tl -> GTypTuple (List.map (fun t -> subst_type subst t) tl) | GTypRef (qid, args) -> GTypRef (qid, List.map (fun t -> subst_type subst t) args) @@ -237,7 +233,8 @@ let rec intern_type env (t : raw_typexpr) : UF.elt glb_typexpr = match t with | RelId (loc, qid) -> let (dp, id) = repr_qualid qid in if DirPath.is_empty dp && Id.Map.mem id env.env_rec then - Id.Map.find id env.env_rec + let (kn, n) = Id.Map.find id env.env_rec in + (Other kn, n) else let kn = try Tac2env.locate_type qid @@ -245,17 +242,20 @@ let rec intern_type env (t : raw_typexpr) : UF.elt glb_typexpr = match t with user_err ?loc (str "Unbound type constructor " ++ pr_qualid qid) in let (nparams, _) = Tac2env.interp_type kn in - (kn, nparams) - | AbsKn kn -> + (Other kn, nparams) + | AbsKn (Other kn) -> let (nparams, _) = Tac2env.interp_type kn in - (kn, nparams) + (Other kn, nparams) + | AbsKn (Tuple n) -> + (Tuple n, n) in let nargs = List.length args in let () = if not (Int.equal nparams nargs) then let loc, qid = match rel with | RelId lid -> lid - | AbsKn kn -> Some loc, shortest_qualid_of_type kn + | AbsKn (Other kn) -> Some loc, shortest_qualid_of_type kn + | AbsKn (Tuple _) -> assert false in user_err ?loc (strbrk "The type constructor " ++ pr_qualid qid ++ strbrk " expects " ++ int nparams ++ strbrk " argument(s), but is here \ @@ -263,7 +263,6 @@ let rec intern_type env (t : raw_typexpr) : UF.elt glb_typexpr = match t with in GTypRef (kn, List.map (fun t -> intern_type env t) args) | CTypArrow (loc, t1, t2) -> GTypArrow (intern_type env t1, intern_type env t2) -| CTypTuple (loc, tl) -> GTypTuple (List.map (fun t -> intern_type env t) tl) let fresh_type_scheme env (t : type_scheme) : UF.elt glb_typexpr = let (n, t) = t in @@ -280,8 +279,11 @@ let fresh_mix_type_scheme env (t : mix_type_scheme) : UF.elt glb_typexpr = in subst_type substf t -let fresh_reftype env (kn : KerName.t) = - let (n, _) = Tac2env.interp_type kn in +let fresh_reftype env (kn : KerName.t or_tuple) = + let n = match kn with + | Other kn -> fst (Tac2env.interp_type kn) + | Tuple n -> n + in let subst = Array.init n (fun _ -> fresh_id env) in let t = GTypRef (kn, Array.map_to_list (fun i -> GTypVar i) subst) in (subst, t) @@ -310,9 +312,9 @@ let rec kind env t = match t with | None -> GTypVar id | Some t -> kind env t end -| GTypRef (kn, tl) -> +| GTypRef (Other kn, tl) -> if is_unfoldable kn then kind env (unfold env kn tl) else t -| GTypArrow _ | GTypTuple _ -> t +| GTypArrow _ | GTypRef (Tuple _, _) -> t exception Occur @@ -321,8 +323,6 @@ let rec occur_check env id t = match kind env t with | GTypArrow (t1, t2) -> let () = occur_check env id t1 in occur_check env id t2 -| GTypTuple tl -> - List.iter (fun t -> occur_check env id t) tl | GTypRef (kn, tl) -> List.iter (fun t -> occur_check env id t) tl @@ -331,24 +331,25 @@ exception CannotUnify of UF.elt glb_typexpr * UF.elt glb_typexpr let unify_var env id t = match kind env t with | GTypVar id' -> if not (UF.equal id id') then UF.union id id' env.env_cst -| GTypArrow _ | GTypRef _ | GTypTuple _ -> +| GTypArrow _ | GTypRef _ -> try let () = occur_check env id t in UF.set id t env.env_cst with Occur -> raise (CannotUnify (GTypVar id, t)) +let eq_or_tuple eq t1 t2 = match t1, t2 with +| Tuple n1, Tuple n2 -> Int.equal n1 n2 +| Other o1, Other o2 -> eq o1 o2 +| _ -> false + let rec unify env t1 t2 = match kind env t1, kind env t2 with | GTypVar id, t | t, GTypVar id -> unify_var env id t | GTypArrow (t1, u1), GTypArrow (t2, u2) -> let () = unify env t1 t2 in unify env u1 u2 -| GTypTuple tl1, GTypTuple tl2 -> - if Int.equal (List.length tl1) (List.length tl2) then - List.iter2 (fun t1 t2 -> unify env t1 t2) tl1 tl2 - else raise (CannotUnify (t1, t2)) | GTypRef (kn1, tl1), GTypRef (kn2, tl2) -> - if KerName.equal kn1 kn2 then + if eq_or_tuple KerName.equal kn1 kn2 then List.iter2 (fun t1 t2 -> unify env t1 t2) tl1 tl2 else raise (CannotUnify (t1, t2)) | _ -> raise (CannotUnify (t1, t2)) @@ -371,7 +372,7 @@ let unify_arrow ?loc env ft args = let ft = GTypVar (fresh_id env) in let () = unify_var env id (GTypArrow (t, ft)) in iter ft args true - | (GTypRef _ | GTypTuple _), _ :: _ -> + | GTypRef _, _ :: _ -> let name = env_name env in if is_fun then user_err ?loc (str "This function has type " ++ pr_glbtype name ft0 ++ @@ -395,10 +396,10 @@ let is_pure_constructor kn = let rec is_value = function | GTacAtm (AtmInt _) | GTacVar _ | GTacRef _ | GTacFun _ -> true | GTacAtm (AtmStr _) | GTacApp _ | GTacLet _ -> false -| GTacCst (GCaseTuple _, _, el) -> List.for_all is_value el +| GTacCst (Tuple _, _, el) -> List.for_all is_value el | GTacCst (_, _, []) -> true | GTacOpn (_, el) -> List.for_all is_value el -| GTacCst (GCaseAlg kn, _, el) -> is_pure_constructor kn && List.for_all is_value el +| GTacCst (Other kn, _, el) -> is_pure_constructor kn && List.for_all is_value el | GTacArr _ | GTacCse _ | GTacPrj _ | GTacSet _ | GTacExt _ | GTacPrm _ | GTacWth _ -> false @@ -411,7 +412,6 @@ let is_rec_rhs = function let rec fv_type f t accu = match t with | GTypVar id -> f id accu | GTypArrow (t1, t2) -> fv_type f t1 (fv_type f t2 accu) -| GTypTuple tl -> List.fold_left (fun accu t -> fv_type f t accu) accu tl | GTypRef (kn, tl) -> List.fold_left (fun accu t -> fv_type f t accu) accu tl let fv_env env = @@ -468,18 +468,19 @@ let warn_redundant_clause = let check_elt_unit loc env t = let maybe_unit = match kind env t with | GTypVar _ -> true - | GTypArrow _ | GTypTuple _ -> false - | GTypRef (kn, _) -> KerName.equal kn t_unit + | GTypArrow _ -> false + | GTypRef (Tuple 0, []) -> true + | GTypRef _ -> false in if not maybe_unit then warn_not_unit ~loc () let check_elt_empty loc env t = match kind env t with | GTypVar _ -> user_err ~loc (str "Cannot infer an empty type for this expression") -| GTypArrow _ | GTypTuple _ -> +| GTypArrow _ | GTypRef (Tuple _, _) -> let name = env_name env in user_err ~loc (str "Type " ++ pr_glbtype name t ++ str " is not an empty type") -| GTypRef (kn, _) -> +| GTypRef (Other kn, _) -> let def = Tac2env.interp_type kn in match def with | _, GTydAlg { galg_constructors = [] } -> kn @@ -488,10 +489,14 @@ let check_elt_empty loc env t = match kind env t with user_err ~loc (str "Type " ++ pr_glbtype name t ++ str " is not an empty type") let check_unit ?loc t = - let maybe_unit = match t with + let env = empty_env () in + (** Should not matter, t should be closed. *) + let t = fresh_type_scheme env t in + let maybe_unit = match kind env t with | GTypVar _ -> true - | GTypArrow _ | GTypTuple _ -> false - | GTypRef (kn, _) -> KerName.equal kn t_unit + | GTypArrow _ -> false + | GTypRef (Tuple 0, []) -> true + | GTypRef _ -> false in if not maybe_unit then warn_not_unit ?loc () @@ -520,15 +525,11 @@ let get_constructor env var = match var with | RelId (loc, qid) -> let c = try Some (Tac2env.locate_constructor qid) with Not_found -> None in begin match c with - | Some knc -> - let kn = Tac2env.interp_constructor knc in - (kn, knc) + | Some knc -> Other knc | None -> CErrors.user_err ?loc (str "Unbound constructor " ++ pr_qualid qid) end -| AbsKn knc -> - let kn = Tac2env.interp_constructor knc in - (kn, knc) +| AbsKn knc -> knc let get_projection var = match var with | RelId (loc, qid) -> @@ -540,37 +541,33 @@ let get_projection var = match var with Tac2env.interp_projection kn let intern_atm env = function -| AtmInt n -> (GTacAtm (AtmInt n), GTypRef (t_int, [])) -| AtmStr s -> (GTacAtm (AtmStr s), GTypRef (t_string, [])) +| AtmInt n -> (GTacAtm (AtmInt n), GTypRef (Other t_int, [])) +| AtmStr s -> (GTacAtm (AtmStr s), GTypRef (Other t_string, [])) -let invalid_pattern ?loc kn t = - let pt = match t with - | GCaseAlg kn' -> pr_typref kn - | GCaseTuple n -> str "tuple" +let invalid_pattern ?loc kn kn' = + let pr t = match t with + | Other kn' -> str "type " ++ pr_typref kn' + | Tuple n -> str "tuple of size " ++ int n in - user_err ?loc (str "Invalid pattern, expected a pattern for type " ++ - pr_typref kn ++ str ", found a pattern of type " ++ pt) (** FIXME *) + user_err ?loc (str "Invalid pattern, expected a pattern for " ++ + pr kn ++ str ", found a pattern for " ++ pr kn') (** FIXME *) (** Pattern view *) type glb_patexpr = | GPatVar of Name.t -| GPatRef of ltac_constructor * glb_patexpr list -| GPatTup of glb_patexpr list +| GPatRef of ltac_constructor or_tuple * glb_patexpr list let rec intern_patexpr env = function | CPatVar (_, na) -> GPatVar na | CPatRef (_, qid, pl) -> - let (_, kn) = get_constructor env qid in + let kn = get_constructor env qid in GPatRef (kn, List.map (fun p -> intern_patexpr env p) pl) -| CPatTup (_, pl) -> - GPatTup (List.map (fun p -> intern_patexpr env p) pl) type pattern_kind = | PKind_empty -| PKind_variant of type_constant +| PKind_variant of type_constant or_tuple | PKind_open of type_constant -| PKind_tuple of int | PKind_any let get_pattern_kind env pl = match pl with @@ -582,11 +579,11 @@ let get_pattern_kind env pl = match pl with | [] -> PKind_any | p :: pl -> get_kind p pl end - | GPatRef (kn, pl) -> + | GPatRef (Other kn, pl) -> let data = Tac2env.interp_constructor kn in if Option.is_empty data.cdata_indx then PKind_open data.cdata_type - else PKind_variant data.cdata_type - | GPatTup tp -> PKind_tuple (List.length tp) + else PKind_variant (Other data.cdata_type) + | GPatRef (Tuple _, tp) -> PKind_variant (Tuple (List.length tp)) in get_kind p pl @@ -611,9 +608,8 @@ let rec intern_rec env = function let (_, _, sch) = Tac2env.interp_global kn in (GTacRef kn, fresh_type_scheme env sch) end -| CTacCst qid as e -> - let loc = loc_of_tacexpr e in - let (_, kn) = get_constructor env qid in +| CTacCst (loc, qid) -> + let kn = get_constructor env qid in intern_constructor env loc kn [] | CTacFun (loc, bnd, e) -> (** Expand pattern: [fun p => t] becomes [fun x => match x with p => t end] *) @@ -646,9 +642,8 @@ let rec intern_rec env = function let (e, t) = intern_rec env e in let t = List.fold_left (fun accu t -> GTypArrow (t, accu)) t tl in (GTacFun (nas, e), t) -| CTacApp (loc, CTacCst qid, args) as e -> - let (_, kn) = get_constructor env qid in - let loc = loc_of_tacexpr e in +| CTacApp (loc, CTacCst (_, qid), args) -> + let kn = get_constructor env qid in intern_constructor env loc kn args | CTacApp (loc, f, args) -> let loc = loc_of_tacexpr f in @@ -688,31 +683,22 @@ let rec intern_rec env = function (GTacLet (false, el, e), t) | CTacLet (loc, true, el, e) -> intern_let_rec env loc el e -| CTacTup (loc, []) -> - (GTacCst (GCaseAlg t_unit, 0, []), GTypRef (t_unit, [])) -| CTacTup (loc, el) -> - let fold e (el, tl) = - let (e, t) = intern_rec env e in - (e :: el, t :: tl) - in - let (el, tl) = List.fold_right fold el ([], []) in - (GTacCst (GCaseTuple (List.length el), 0, el), GTypTuple tl) | CTacArr (loc, []) -> let id = fresh_id env in - (GTacArr [], GTypRef (t_int, [GTypVar id])) + (GTacArr [], GTypRef (Other t_int, [GTypVar id])) | CTacArr (loc, e0 :: el) -> let (e0, t0) = intern_rec env e0 in let fold e el = intern_rec_with_constraint env e t0 :: el in let el = e0 :: List.fold_right fold el [] in - (GTacArr el, GTypRef (t_array, [t0])) + (GTacArr el, GTypRef (Other t_array, [t0])) | CTacLst (loc, []) -> let id = fresh_id env in - (c_nil, GTypRef (t_list, [GTypVar id])) + (c_nil, GTypRef (Other t_list, [GTypVar id])) | CTacLst (loc, e0 :: el) -> let (e0, t0) = intern_rec env e0 in let fold e el = c_cons (intern_rec_with_constraint env e t0) el in let el = c_cons e0 (List.fold_right fold el c_nil) in - (el, GTypRef (t_list, [t0])) + (el, GTypRef (Other t_list, [t0])) | CTacCnv (loc, e, tc) -> let (e, t) = intern_rec env e in let tc = intern_type env tc in @@ -733,7 +719,7 @@ let rec intern_rec env = function let (e, t) = intern_rec env e in let subst = Array.init pinfo.pdata_prms (fun _ -> fresh_id env) in let params = Array.map_to_list (fun i -> GTypVar i) subst in - let exp = GTypRef (pinfo.pdata_type, params) in + let exp = GTypRef (Other pinfo.pdata_type, params) in let () = unify ~loc env t exp in let substf i = GTypVar subst.(i) in let ret = subst_type substf pinfo.pdata_ptyp in @@ -750,12 +736,12 @@ let rec intern_rec env = function in let subst = Array.init pinfo.pdata_prms (fun _ -> fresh_id env) in let params = Array.map_to_list (fun i -> GTypVar i) subst in - let exp = GTypRef (pinfo.pdata_type, params) in + let exp = GTypRef (Other pinfo.pdata_type, params) in let e = intern_rec_with_constraint env e exp in let substf i = GTypVar subst.(i) in let ret = subst_type substf pinfo.pdata_ptyp in let r = intern_rec_with_constraint env r ret in - (GTacSet (pinfo.pdata_type, e, pinfo.pdata_indx, r), GTypRef (t_unit, [])) + (GTacSet (pinfo.pdata_type, e, pinfo.pdata_indx, r), GTypRef (Tuple 0, [])) | CTacExt (loc, ext) -> let open Genintern in let GenArg (Rawwit tag, _) = ext in @@ -766,7 +752,7 @@ let rec intern_rec env = function let ist = empty_glob_sign genv in let ist = { ist with extra = Store.set ist.extra ltac2_env env } in let (_, ext) = Flags.with_option Ltac_plugin.Tacintern.strict_check (fun () -> generic_intern ist ext) () in - (GTacExt ext, GTypRef (tpe.ml_type, [])) + (GTacExt ext, GTypRef (Other tpe.ml_type, [])) and intern_rec_with_constraint env e exp = let loc = loc_of_tacexpr e in @@ -830,39 +816,21 @@ and intern_case env loc e pl = | PKind_empty -> let kn = check_elt_empty loc env t in let r = fresh_id env in - (GTacCse (e', GCaseAlg kn, [||], [||]), GTypVar r) - | PKind_tuple len -> - begin match pl with - | [] -> assert false - | [CPatTup (_, []), b] -> - let () = unify ~loc:(loc_of_tacexpr e) env t (GTypRef (t_unit, [])) in - let (b, tb) = intern_rec env b in - (GTacCse (e', GCaseAlg t_unit, [|b|], [||]), tb) - | [CPatTup (_, pl), b] -> - let map = function - | CPatVar (_, na) -> na - | p -> todo ~loc:(loc_of_patexpr p) () - in - let ids = List.map map pl in - let targs = List.map (fun _ -> GTypVar (fresh_id env)) pl in - let tc = GTypTuple targs in - let () = unify ~loc:(loc_of_tacexpr e) env t tc in - let env = List.fold_left2 (fun env na t -> push_name na (monomorphic t) env) env ids targs in - let (b, tb) = intern_rec env b in - (GTacCse (e', GCaseTuple len, [||], [|Array.of_list ids, b|]), tb) - | (p, _) :: _ -> todo ~loc:(loc_of_patexpr p) () - end + (GTacCse (e', Other kn, [||], [||]), GTypVar r) | PKind_variant kn -> let subst, tc = fresh_reftype env kn in let () = unify ~loc:(loc_of_tacexpr e) env t tc in - let (params, def) = Tac2env.interp_type kn in - let galg = match def with - | GTydAlg c -> c - | _ -> assert false + let (nconst, nnonconst, arities) = match kn with + | Tuple 0 -> 1, 0, [0] + | Tuple n -> 0, 1, [n] + | Other kn -> + let (_, def) = Tac2env.interp_type kn in + let galg = match def with | GTydAlg c -> c | _ -> assert false in + let arities = List.map (fun (_, args) -> List.length args) galg.galg_constructors in + galg.galg_nconst, galg.galg_nnonconst, arities in - let cstrs = galg.galg_constructors in - let const = Array.make galg.galg_nconst None in - let nonconst = Array.make galg.galg_nnonconst None in + let const = Array.make nconst None in + let nonconst = Array.make nnonconst None in let ret = GTypVar (fresh_id env) in let rec intern_branch = function | [] -> () @@ -873,8 +841,8 @@ and intern_case env loc e pl = let () = check_redundant_clause rem in let (br', brT) = intern_rec env br in (** Fill all remaining branches *) - let fill (ncst, narg) (_, args) = - if List.is_empty args then + let fill (ncst, narg) arity = + if Int.equal arity 0 then let () = if Option.is_empty const.(ncst) then const.(ncst) <- Some br' in @@ -882,19 +850,25 @@ and intern_case env loc e pl = else let () = if Option.is_empty nonconst.(narg) then - let ids = Array.map_of_list (fun _ -> Anonymous) args in + let ids = Array.make arity Anonymous in nonconst.(narg) <- Some (ids, br') in (ncst, succ narg) in - let _ = List.fold_left fill (0, 0) cstrs in + let _ = List.fold_left fill (0, 0) arities in brT | CPatRef (loc, qid, args) -> - let (data, _) = get_constructor env qid in + let knc = get_constructor env qid in + let kn', index, arity = match knc with + | Tuple n -> Tuple n, 0, List.init n (fun i -> GTypVar i) + | Other knc -> + let data = Tac2env.interp_constructor knc in + let index = Option.get data.cdata_indx in + Other data.cdata_type, index, data.cdata_args + in let () = - let kn' = data.cdata_type in - if not (KerName.equal kn kn') then - invalid_pattern ~loc kn (GCaseAlg kn') + if not (eq_or_tuple KerName.equal kn kn') then + invalid_pattern ~loc kn kn' in let get_id = function | CPatVar (_, na) -> na @@ -902,7 +876,7 @@ and intern_case env loc e pl = in let ids = List.map get_id args in let nids = List.length ids in - let nargs = List.length data.cdata_args in + let nargs = List.length arity in let () = if not (Int.equal nids nargs) then error_nargs_mismatch loc nargs nids in @@ -912,13 +886,9 @@ and intern_case env loc e pl = let tpe = subst_type subst tpe in push_name id (monomorphic tpe) env in - let nenv = List.fold_left2 fold env ids data.cdata_args in + let nenv = List.fold_left2 fold env ids arity in let (br', brT) = intern_rec nenv br in let () = - let index = match data.cdata_indx with - | Some i -> i - | None -> assert false - in if List.is_empty args then if Option.is_empty const.(index) then const.(index) <- Some br' else warn_redundant_clause ~loc () @@ -928,8 +898,6 @@ and intern_case env loc e pl = else warn_redundant_clause ~loc () in brT - | CPatTup (loc, tup) -> - invalid_pattern ?loc kn (GCaseTuple (List.length tup)) in let () = unify ~loc:(loc_of_tacexpr br) env ret tbr in intern_branch rem @@ -941,10 +909,10 @@ and intern_case env loc e pl = in let const = Array.map map const in let nonconst = Array.map map nonconst in - let ce = GTacCse (e', GCaseAlg kn, const, nonconst) in + let ce = GTacCse (e', kn, const, nonconst) in (ce, ret) | PKind_open kn -> - let subst, tc = fresh_reftype env kn in + let subst, tc = fresh_reftype env (Other kn) in let () = unify ~loc:(loc_of_tacexpr e) env t tc in let ret = GTypVar (fresh_id env) in let rec intern_branch map = function @@ -961,15 +929,19 @@ and intern_case env loc e pl = | GPatRef (knc, args) -> let get = function | GPatVar na -> na - | GPatRef _ | GPatTup _ -> + | GPatRef _ -> user_err ~loc (str "TODO: Unhandled match case") (** FIXME *) in let loc = loc_of_patexpr pat in + let knc = match knc with + | Other knc -> knc + | Tuple n -> invalid_pattern ~loc (Other kn) (Tuple n) + in let ids = List.map get args in let data = Tac2env.interp_constructor knc in let () = if not (KerName.equal kn data.cdata_type) then - invalid_pattern ~loc kn (GCaseAlg data.cdata_type) + invalid_pattern ~loc (Other kn) (Other data.cdata_type) in let nids = List.length ids in let nargs = List.length data.cdata_args in @@ -992,29 +964,36 @@ and intern_case env loc e pl = KNmap.add knc (Anonymous, Array.of_list ids, br') map in intern_branch map rem - | GPatTup tup -> - invalid_pattern ~loc kn (GCaseTuple (List.length tup)) in let (map, def) = intern_branch KNmap.empty pl in (GTacWth { opn_match = e'; opn_branch = map; opn_default = def }, ret) -and intern_constructor env loc kn args = +and intern_constructor env loc kn args = match kn with +| Other kn -> let cstr = interp_constructor kn in let nargs = List.length cstr.cdata_args in if Int.equal nargs (List.length args) then let subst = Array.init cstr.cdata_prms (fun _ -> fresh_id env) in let substf i = GTypVar subst.(i) in let types = List.map (fun t -> subst_type substf t) cstr.cdata_args in - let ans = GTypRef (cstr.cdata_type, List.init cstr.cdata_prms (fun i -> GTypVar subst.(i))) in + let targs = List.init cstr.cdata_prms (fun i -> GTypVar subst.(i)) in + let ans = GTypRef (Other cstr.cdata_type, targs) in let map arg tpe = intern_rec_with_constraint env arg tpe in let args = List.map2 map args types in match cstr.cdata_indx with | Some idx -> - (GTacCst (GCaseAlg cstr.cdata_type, idx, args), ans) + (GTacCst (Other cstr.cdata_type, idx, args), ans) | None -> (GTacOpn (kn, args), ans) else error_nargs_mismatch loc nargs (List.length args) +| Tuple n -> + assert (Int.equal n (List.length args)); + let types = List.init n (fun i -> GTypVar (fresh_id env)) in + let map arg tpe = intern_rec_with_constraint env arg tpe in + let args = List.map2 map args types in + let ans = GTypRef (Tuple n, types) in + GTacCst (Tuple n, 0, args), ans and intern_record env loc fs = let map (proj, e) = @@ -1062,7 +1041,7 @@ and intern_record env loc fs = in let args = Array.map_to_list Option.get args in let tparam = List.init params (fun i -> GTypVar subst.(i)) in - (GTacCst (GCaseAlg kn, 0, args), GTypRef (kn, tparam)) + (GTacCst (Other kn, 0, args), GTypRef (Other kn, tparam)) let normalize env (count, vars) (t : UF.elt glb_typexpr) = let get_var id = @@ -1154,8 +1133,6 @@ let rec ids_of_pattern accu = function | CPatVar (_, Name id) -> Id.Set.add id accu | CPatRef (_, _, pl) -> List.fold_left ids_of_pattern accu pl -| CPatTup (_, pl) -> - List.fold_left ids_of_pattern accu pl let rec globalize ids e = match e with | CTacAtm _ -> e @@ -1165,9 +1142,9 @@ let rec globalize ids e = match e with | ArgVar _ -> e | ArgArg kn -> CTacRef (AbsKn kn) end -| CTacCst qid -> - let (_, knc) = get_constructor () qid in - CTacCst (AbsKn knc) +| CTacCst (loc, qid) -> + let knc = get_constructor () qid in + CTacCst (loc, AbsKn knc) | CTacFun (loc, bnd, e) -> let fold (pats, accu) (pat, t) = let accu = ids_of_pattern accu pat in @@ -1193,9 +1170,6 @@ let rec globalize ids e = match e with in let bnd = List.map map bnd in CTacLet (loc, isrec, bnd, e) -| CTacTup (loc, el) -> - let el = List.map (fun e -> globalize ids e) el in - CTacTup (loc, el) | CTacArr (loc, el) -> let el = List.map (fun e -> globalize ids e) el in CTacArr (loc, el) @@ -1239,18 +1213,21 @@ and globalize_case ids (p, e) = and globalize_pattern ids p = match p with | CPatVar _ -> p | CPatRef (loc, cst, pl) -> - let (_, knc) = get_constructor () cst in + let knc = get_constructor () cst in let cst = AbsKn knc in let pl = List.map (fun p -> globalize_pattern ids p) pl in CPatRef (loc, cst, pl) -| CPatTup (loc, pl) -> - let pl = List.map (fun p -> globalize_pattern ids p) pl in - CPatTup (loc, pl) (** Kernel substitution *) open Mod_subst +let subst_or_tuple f subst o = match o with +| Tuple _ -> o +| Other v -> + let v' = f subst v in + if v' == v then o else Other v' + let rec subst_type subst t = match t with | GTypVar _ -> t | GTypArrow (t1, t2) -> @@ -1258,20 +1235,11 @@ let rec subst_type subst t = match t with let t2' = subst_type subst t2 in if t1' == t1 && t2' == t2 then t else GTypArrow (t1', t2') -| GTypTuple tl -> - let tl'= List.smartmap (fun t -> subst_type subst t) tl in - if tl' == tl then t else GTypTuple tl' | GTypRef (kn, tl) -> - let kn' = subst_kn subst kn in + let kn' = subst_or_tuple subst_kn subst kn in let tl' = List.smartmap (fun t -> subst_type subst t) tl in if kn' == kn && tl' == tl then t else GTypRef (kn', tl') -let subst_case_info subst ci = match ci with -| GCaseAlg kn -> - let kn' = subst_kn subst kn in - if kn' == kn then ci else GCaseAlg kn' -| GCaseTuple _ -> ci - let rec subst_expr subst e = match e with | GTacAtm _ | GTacVar _ | GTacPrm _ -> e | GTacRef kn -> GTacRef (subst_kn subst kn) @@ -1284,18 +1252,13 @@ let rec subst_expr subst e = match e with | GTacArr el -> GTacArr (List.map (fun e -> subst_expr subst e) el) | GTacCst (t, n, el) as e0 -> - let t' = match t with - | GCaseAlg kn -> - let kn' = subst_kn subst kn in - if kn' == kn then t else GCaseAlg kn' - | GCaseTuple _ -> t - in + let t' = subst_or_tuple subst_kn subst t in let el' = List.smartmap (fun e -> subst_expr subst e) el in if t' == t && el' == el then e0 else GTacCst (t', n, el') | GTacCse (e, ci, cse0, cse1) -> let cse0' = Array.map (fun e -> subst_expr subst e) cse0 in let cse1' = Array.map (fun (ids, e) -> (ids, subst_expr subst e)) cse1 in - let ci' = subst_case_info subst ci in + let ci' = subst_or_tuple subst_kn subst ci in GTacCse (subst_expr subst e, ci', cse0', cse1') | GTacWth { opn_match = e; opn_branch = br; opn_default = (na, def) } as e0 -> let e' = subst_expr subst e in @@ -1358,7 +1321,7 @@ let subst_type_scheme subst (prm, t as sch) = let subst_or_relid subst ref = match ref with | RelId _ -> ref | AbsKn kn -> - let kn' = subst_kn subst kn in + let kn' = subst_or_tuple subst_kn subst kn in if kn' == kn then ref else AbsKn kn' let rec subst_rawtype subst t = match t with @@ -1367,9 +1330,6 @@ let rec subst_rawtype subst t = match t with let t1' = subst_rawtype subst t1 in let t2' = subst_rawtype subst t2 in if t1' == t1 && t2' == t2 then t else CTypArrow (loc, t1', t2') -| CTypTuple (loc, tl) -> - let tl' = List.smartmap (fun t -> subst_rawtype subst t) tl in - if tl' == tl then t else CTypTuple (loc, tl') | CTypRef (loc, ref, tl) -> let ref' = subst_or_relid subst ref in let tl' = List.smartmap (fun t -> subst_rawtype subst t) tl in @@ -1391,16 +1351,8 @@ let rec subst_rawpattern subst p = match p with | CPatVar _ -> p | CPatRef (loc, c, pl) -> let pl' = List.smartmap (fun p -> subst_rawpattern subst p) pl in - let c' = match c with - | RelId _ -> c - | AbsKn kn -> - let kn' = subst_kn subst kn in - if kn' == kn then c else AbsKn kn' - in + let c' = subst_or_relid subst c in if pl' == pl && c' == c then p else CPatRef (loc, c', pl') -| CPatTup (loc, pl) -> - let pl' = List.smartmap (fun p -> subst_rawpattern subst p) pl in - if pl' == pl then p else CPatTup (loc, pl') (** Used for notations *) let rec subst_rawexpr subst t = match t with @@ -1408,9 +1360,9 @@ let rec subst_rawexpr subst t = match t with | CTacRef ref -> let ref' = subst_tacref subst ref in if ref' == ref then t else CTacRef ref' -| CTacCst ref -> - let ref' = subst_tacref subst ref in - if ref' == ref then t else CTacCst ref' +| CTacCst (loc, ref) -> + let ref' = subst_or_relid subst ref in + if ref' == ref then t else CTacCst (loc, ref') | CTacFun (loc, bnd, e) -> let map (na, t as p) = let t' = Option.smartmap (fun t -> subst_rawtype subst t) t in @@ -1432,9 +1384,6 @@ let rec subst_rawexpr subst t = match t with let bnd' = List.smartmap map bnd in let e' = subst_rawexpr subst e in if bnd' == bnd && e' == e then t else CTacLet (loc, isrec, bnd', e') -| CTacTup (loc, el) -> - let el' = List.smartmap (fun e -> subst_rawexpr subst e) el in - if el' == el then t else CTacTup (loc, el') | CTacArr (loc, el) -> let el' = List.smartmap (fun e -> subst_rawexpr subst e) el in if el' == el then t else CTacArr (loc, el') diff --git a/src/tac2intern.mli b/src/tac2intern.mli index b2604c4ea7..ddec8eb7e4 100644 --- a/src/tac2intern.mli +++ b/src/tac2intern.mli @@ -21,7 +21,7 @@ val intern_open_type : raw_typexpr -> type_scheme (** Check that a term is a value. Only values are safe to marshall between processes. *) val is_value : glb_tacexpr -> bool -val check_unit : ?loc:Loc.t -> int glb_typexpr -> unit +val check_unit : ?loc:Loc.t -> type_scheme -> unit val subst_type : substitution -> 'a glb_typexpr -> 'a glb_typexpr val subst_expr : substitution -> glb_tacexpr -> glb_tacexpr diff --git a/src/tac2print.ml b/src/tac2print.ml index aa8e1e98d5..28f9516f65 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -30,20 +30,23 @@ type typ_level = | T1 | T0 +let t_unit = + KerName.make2 Tac2env.coq_prefix (Label.of_id (Id.of_string "unit")) + let pr_typref kn = Libnames.pr_qualid (Tac2env.shortest_qualid_of_type kn) let pr_glbtype_gen pr lvl c = let rec pr_glbtype lvl = function | GTypVar n -> str "'" ++ str (pr n) - | GTypRef (kn, []) -> pr_typref kn - | GTypRef (kn, [t]) -> + | GTypRef (Other kn, []) -> pr_typref kn + | GTypRef (Other kn, [t]) -> let paren = match lvl with | T5_r | T5_l | T2 | T1 -> fun x -> x | T0 -> paren in paren (pr_glbtype lvl t ++ spc () ++ pr_typref kn) - | GTypRef (kn, tl) -> + | GTypRef (Other kn, tl) -> let paren = match lvl with | T5_r | T5_l | T2 | T1 -> fun x -> x | T0 -> paren @@ -55,7 +58,9 @@ let pr_glbtype_gen pr lvl c = | T5_l | T2 | T1 | T0 -> paren in paren (pr_glbtype T5_l t1 ++ spc () ++ str "->" ++ spc () ++ pr_glbtype T5_r t2) - | GTypTuple tl -> + | GTypRef (Tuple 0, []) -> + Libnames.pr_qualid (Tac2env.shortest_qualid_of_type t_unit) + | GTypRef (Tuple _, tl) -> let paren = match lvl with | T5_r | T5_l -> fun x -> x | T2 | T1 | T0 -> paren @@ -165,7 +170,8 @@ let pr_glbexpr_gen lvl c = in let bnd = prlist_with_sep (fun () -> str "with" ++ spc ()) pr_bnd bnd in paren (str "let" ++ spc () ++ mut ++ bnd ++ str "in" ++ spc () ++ pr_glbexpr E5 e) - | GTacCst (GCaseTuple _, _, cl) -> + | GTacCst (Tuple 0, _, _) -> str "()" + | GTacCst (Tuple _, _, cl) -> let paren = match lvl with | E0 | E1 -> paren | E2 | E3 | E4 | E5 -> fun x -> x @@ -173,7 +179,7 @@ let pr_glbexpr_gen lvl c = paren (prlist_with_sep (fun () -> str "," ++ spc ()) (pr_glbexpr E1) cl) | GTacArr cl -> mt () (** FIXME when implemented *) - | GTacCst (GCaseAlg tpe, n, cl) -> + | GTacCst (Other tpe, n, cl) -> begin match Tac2env.interp_type tpe with | _, GTydAlg { galg_constructors = def } -> let paren = match lvl with @@ -200,7 +206,7 @@ let pr_glbexpr_gen lvl c = | GTacCse (e, info, cst_br, ncst_br) -> let e = pr_glbexpr E5 e in let br = match info with - | GCaseAlg kn -> + | Other kn -> let def = match Tac2env.interp_type kn with | _, GTydAlg { galg_constructors = def } -> def | _, GTydDef _ | _, GTydRec _ | _, GTydOpn -> assert false @@ -217,8 +223,8 @@ let pr_glbexpr_gen lvl c = hov 2 (pr_glbexpr E5 p)) ++ spc () in prlist pr_branch br - | GCaseTuple n -> - let (vars, p) = ncst_br.(0) in + | Tuple n -> + let (vars, p) = if Int.equal n 0 then ([||], cst_br.(0)) else ncst_br.(0) in let p = pr_glbexpr E5 p in let vars = prvect_with_sep (fun () -> str "," ++ spc ()) pr_name vars in str "|" ++ spc () ++ paren vars ++ spc () ++ str "=>" ++ spc () ++ p diff --git a/tests/stuff/ltac2.v b/tests/stuff/ltac2.v index 86ab5afb17..4950a20ec4 100644 --- a/tests/stuff/ltac2.v +++ b/tests/stuff/ltac2.v @@ -120,6 +120,7 @@ Abort. Goal True. Proof. + let x () := plus (fun () => 0) (fun _ => 1) in match case x with | Val x => -- cgit v1.2.3 From 4b863ed5a4c9545ecfd25dc22a83edd3c47aab80 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 28 Jul 2017 15:23:16 +0200 Subject: Allowing generic patterns in let-bindings. --- src/g_ltac2.ml4 | 18 ++++-- src/tac2entries.ml | 4 +- src/tac2expr.mli | 2 +- src/tac2intern.ml | 170 +++++++++++++++++++++++++++++------------------------ 4 files changed, 111 insertions(+), 83 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 7ee9d7e282..21612f9a25 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -117,10 +117,20 @@ GEXTEND Gram ] ] ; let_clause: - [ [ id = binder; ":="; te = tac2expr -> - (id, None, te) - | id = binder; args = LIST1 input_fun; ":="; te = tac2expr -> - (id, None, CTacFun (!@loc, args, te)) ] ] + [ [ binder = let_binder; ":="; te = tac2expr -> + let (pat, fn) = binder in + let te = match fn with None -> te | Some args -> CTacFun (!@loc, args, te) in + (pat, None, te) + ] ] + ; + let_binder: + [ [ pats = LIST1 input_fun -> + match pats with + | [CPatVar _ as pat, None] -> (pat, None) + | (CPatVar (_, Name id) as pat, None) :: args -> (pat, Some args) + | [pat, None] -> (pat, None) + | _ -> CErrors.user_err ~loc:!@loc (str "Invalid pattern") + ] ] ; tac2type: [ "5" RIGHTA diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 70f1b583e0..7bc4c75510 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -266,7 +266,7 @@ let inline_rec_tactic tactics = in (** Fresh variables to abstract over the function patterns *) let _, vars = List.fold_left fold_var (avoid, []) pat in - let map_body ((loc, id), _, e) = (loc, Name id), None, e in + let map_body ((loc, id), _, e) = CPatVar (loc, Name id), None, e in let bnd = List.map map_body tactics in let pat_of_id (loc, id) = (CPatVar (loc, Name id), None) @@ -552,7 +552,7 @@ let perform_notation syn st = let mk loc args = let map (na, e) = let loc = loc_of_tacexpr e in - (Loc.tag ~loc na, None, e) + (CPatVar (Loc.tag ~loc na), None, e) in let bnd = List.map map args in CTacLet (loc, false, bnd, syn.synext_exp) diff --git a/src/tac2expr.mli b/src/tac2expr.mli index 668bc0dad1..10d8c1d421 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -92,7 +92,7 @@ type raw_tacexpr = | CTacCst of Loc.t * ltac_constructor or_tuple or_relid | CTacFun of Loc.t * (raw_patexpr * raw_typexpr option) list * raw_tacexpr | CTacApp of Loc.t * raw_tacexpr * raw_tacexpr list -| CTacLet of Loc.t * rec_flag * (Name.t located * raw_typexpr option * raw_tacexpr) list * raw_tacexpr +| CTacLet of Loc.t * rec_flag * (raw_patexpr * raw_typexpr option * raw_tacexpr) list * raw_tacexpr | CTacArr of raw_tacexpr list located | CTacLst of raw_tacexpr list located | CTacCnv of Loc.t * raw_tacexpr * raw_typexpr diff --git a/src/tac2intern.ml b/src/tac2intern.ml index e460111fc1..16e0bc8cbe 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -590,13 +590,53 @@ let get_pattern_kind env pl = match pl with (** Internalization *) (** Used to generate a fresh tactic variable for pattern-expansion *) -let fresh_var env = +let fresh_var avoid = let bad id = - Id.Map.mem id env.env_var || + Id.Set.mem id avoid || (try ignore (locate_ltac (qualid_of_ident id)); true with Not_found -> false) in Namegen.next_ident_away_from (Id.of_string "p") bad +let add_name accu = function +| Name id -> Id.Set.add id accu +| Anonymous -> accu + +let rec ids_of_pattern accu = function +| CPatVar (_, Anonymous) -> accu +| CPatVar (_, Name id) -> Id.Set.add id accu +| CPatRef (_, _, pl) -> + List.fold_left ids_of_pattern accu pl + +let loc_of_relid = function +| RelId (loc, _) -> Option.default dummy_loc loc +| AbsKn _ -> dummy_loc + +(** Expand pattern: [p => t] becomes [x => match x with p => t end] *) +let expand_pattern avoid bnd = + let fold (avoid, bnd) (pat, t) = + let na, expand = match pat with + | CPatVar (_, na) -> + (** Don't expand variable patterns *) + na, None + | _ -> + let loc = loc_of_patexpr pat in + let id = fresh_var avoid in + let qid = RelId (Loc.tag ~loc (qualid_of_ident id)) in + Name id, Some qid + in + let avoid = ids_of_pattern avoid pat in + let avoid = add_name avoid na in + (avoid, (na, pat, expand) :: bnd) + in + let (_, bnd) = List.fold_left fold (avoid, []) bnd in + let fold e (na, pat, expand) = match expand with + | None -> e + | Some qid -> CTacCse (loc_of_relid qid, CTacRef qid, [pat, e]) + in + let expand e = List.fold_left fold e bnd in + let nas = List.rev_map (fun (na, _, _) -> na) bnd in + (nas, expand) + let rec intern_rec env = function | CTacAtm (_, atm) -> intern_atm env atm | CTacRef qid -> @@ -612,35 +652,15 @@ let rec intern_rec env = function let kn = get_constructor env qid in intern_constructor env loc kn [] | CTacFun (loc, bnd, e) -> - (** Expand pattern: [fun p => t] becomes [fun x => match x with p => t end] *) - let fold (env, bnd, tl) (pat, t) = - let t = match t with - | None -> GTypVar (fresh_id env) - | Some t -> intern_type env t - in - let na, expand = match pat with - | CPatVar (_, na) -> - (** Don't expand variable patterns *) - na, None - | _ -> - let loc = loc_of_patexpr pat in - let id = fresh_var env in - let qid = RelId (Loc.tag ~loc (qualid_of_ident id)) in - Name id, Some qid - in - let env = push_name na (monomorphic t) env in - (env, (na, pat, expand) :: bnd, t :: tl) + let map (_, t) = match t with + | None -> GTypVar (fresh_id env) + | Some t -> intern_type env t in - let (env, bnd, tl) = List.fold_left fold (env, [], []) bnd in - let fold e (na, pat, expand) = match expand with - | None -> e - | Some qid -> - CTacCse (loc, CTacRef qid, [pat, e]) - in - let e = List.fold_left fold e bnd in - let nas = List.rev_map (fun (na, _, _) -> na) bnd in - let (e, t) = intern_rec env e in - let t = List.fold_left (fun accu t -> GTypArrow (t, accu)) t tl in + let tl = List.map map bnd in + let (nas, exp) = expand_pattern (Id.Map.domain env.env_var) bnd in + let env = List.fold_left2 (fun env na t -> push_name na (monomorphic t) env) env nas tl in + let (e, t) = intern_rec env (exp e) in + let t = List.fold_right (fun t accu -> GTypArrow (t, accu)) tl t in (GTacFun (nas, e), t) | CTacApp (loc, CTacCst (_, qid), args) -> let kn = get_constructor env qid in @@ -656,33 +676,20 @@ let rec intern_rec env = function let (args, t) = List.fold_right fold args ([], []) in let ret = unify_arrow ~loc env ft t in (GTacApp (f, args), ret) -| CTacLet (loc, false, el, e) -> - let fold accu ((loc, na), _, _) = match na with - | Anonymous -> accu - | Name id -> - if Id.Set.mem id accu then - user_err ?loc (str "Variable " ++ Id.print id ++ str " is bound several \ +| CTacLet (loc, is_rec, el, e) -> + let fold accu (pat, _, e) = + let ids = ids_of_pattern Id.Set.empty pat in + let common = Id.Set.inter ids accu in + if Id.Set.is_empty common then Id.Set.union ids accu + else + let id = Id.Set.choose common in + let loc = loc_of_patexpr pat in + user_err ~loc (str "Variable " ++ Id.print id ++ str " is bound several \ times in this matching") - else Id.Set.add id accu in - let _ = List.fold_left fold Id.Set.empty el in - let fold ((loc, na), tc, e) (el, p) = - let (e, t) = intern_rec env e in - let () = match tc with - | None -> () - | Some tc -> - let tc = intern_type env tc in - unify ?loc env t tc - in - let t = if is_value e then abstract_var env t else monomorphic t in - ((na, e) :: el), ((na, t) :: p) - in - let (el, p) = List.fold_right fold el ([], []) in - let nenv = List.fold_left (fun accu (na, t) -> push_name na t env) env p in - let (e, t) = intern_rec nenv e in - (GTacLet (false, el, e), t) -| CTacLet (loc, true, el, e) -> - intern_let_rec env loc el e + let ids = List.fold_left fold Id.Set.empty el in + if is_rec then intern_let_rec env loc ids el e + else intern_let env loc ids el e | CTacArr (loc, []) -> let id = fresh_id env in (GTacArr [], GTypRef (Other t_int, [GTypVar id])) @@ -760,17 +767,38 @@ and intern_rec_with_constraint env e exp = let () = unify ~loc env t exp in e -and intern_let_rec env loc el e = - let fold accu ((loc, na), _, _) = match na with - | Anonymous -> accu - | Name id -> - if Id.Set.mem id accu then - user_err ?loc (str "Variable " ++ Id.print id ++ str " is bound several \ - times in this matching") - else Id.Set.add id accu +and intern_let env loc ids el e = + let avoid = Id.Set.union ids (Id.Map.domain env.env_var) in + let fold (pat, t, e) (avoid, accu) = + let nas, exp = expand_pattern avoid [pat, t] in + let na = match nas with [x] -> x | _ -> assert false in + let avoid = List.fold_left add_name avoid nas in + (avoid, (na, exp, t, e) :: accu) in - let _ = List.fold_left fold Id.Set.empty el in - let map env ((loc, na), t, e) = + let (_, el) = List.fold_right fold el (avoid, []) in + let fold (na, exp, tc, e) (body, el, p) = + let (e, t) = match tc with + | None -> intern_rec env e + | Some tc -> + let tc = intern_type env tc in + (intern_rec_with_constraint env e tc, tc) + in + let t = if is_value e then abstract_var env t else monomorphic t in + (exp body, (na, e) :: el, (na, t) :: p) + in + let (e, el, p) = List.fold_right fold el (e, [], []) in + let env = List.fold_left (fun accu (na, t) -> push_name na t accu) env p in + let (e, t) = intern_rec env e in + (GTacLet (false, el, e), t) + +and intern_let_rec env loc ids el e = + let map env (pat, t, e) = + let loc, na = match pat with + | CPatVar na -> na + | CPatRef _ -> + let loc = loc_of_patexpr pat in + user_err ~loc (str "This kind of pattern is forbidden in let-rec bindings") + in let id = fresh_id env in let env = push_name na (monomorphic (GTypVar id)) env in (env, (loc, na, t, e, id)) @@ -1116,10 +1144,6 @@ let intern_open_type t = (** Globalization *) -let add_name accu = function -| Name id -> Id.Set.add id accu -| Anonymous -> accu - let get_projection0 var = match var with | RelId (loc, qid) -> let kn = try Tac2env.locate_projection qid with Not_found -> @@ -1128,12 +1152,6 @@ let get_projection0 var = match var with kn | AbsKn kn -> kn -let rec ids_of_pattern accu = function -| CPatVar (_, Anonymous) -> accu -| CPatVar (_, Name id) -> Id.Set.add id accu -| CPatRef (_, _, pl) -> - List.fold_left ids_of_pattern accu pl - let rec globalize ids e = match e with | CTacAtm _ -> e | CTacRef ref -> @@ -1160,7 +1178,7 @@ let rec globalize ids e = match e with let el = List.map (fun e -> globalize ids e) el in CTacApp (loc, e, el) | CTacLet (loc, isrec, bnd, e) -> - let fold accu ((_, na), _, _) = add_name accu na in + let fold accu (pat, _, _) = ids_of_pattern accu pat in let ext = List.fold_left fold Id.Set.empty bnd in let eids = Id.Set.union ext ids in let e = globalize eids e in -- cgit v1.2.3 From 8aef0199bed6fde2233704deda4116453fca869f Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 28 Jul 2017 18:05:58 +0200 Subject: Moving the Ltac2 FFI to a separate file. --- _CoqProject | 2 + src/ltac2_plugin.mlpack | 1 + src/tac2core.ml | 164 ++++++++---------------------------------------- src/tac2core.mli | 44 ------------- src/tac2ffi.ml | 122 +++++++++++++++++++++++++++++++++++ src/tac2ffi.mli | 78 +++++++++++++++++++++++ src/tac2stdlib.ml | 4 +- 7 files changed, 233 insertions(+), 182 deletions(-) create mode 100644 src/tac2ffi.ml create mode 100644 src/tac2ffi.mli diff --git a/_CoqProject b/_CoqProject index e3ad3987bd..6c9393628d 100644 --- a/_CoqProject +++ b/_CoqProject @@ -13,6 +13,8 @@ src/tac2interp.ml src/tac2interp.mli src/tac2entries.ml src/tac2entries.mli +src/tac2ffi.ml +src/tac2ffi.mli src/tac2core.ml src/tac2core.mli src/tac2stdlib.ml diff --git a/src/ltac2_plugin.mlpack b/src/ltac2_plugin.mlpack index dc78207291..1d7b655dce 100644 --- a/src/ltac2_plugin.mlpack +++ b/src/ltac2_plugin.mlpack @@ -3,6 +3,7 @@ Tac2print Tac2intern Tac2interp Tac2entries +Tac2ffi Tac2core Tac2stdlib G_ltac2 diff --git a/src/tac2core.ml b/src/tac2core.ml index a3678d1ad0..6d9ede4421 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -18,30 +18,9 @@ open Proofview.Notations (** Standard values *) -let coq_core n = KerName.make2 Tac2env.coq_prefix (Label.of_id (Id.of_string_soft n)) +module Value = Tac2ffi -let val_tag t = match val_tag t with -| Val.Base t -> t -| _ -> assert false - -let val_constr = val_tag (topwit Stdarg.wit_constr) -let val_ident = val_tag (topwit Stdarg.wit_ident) -let val_pattern = Val.create "ltac2:pattern" -let val_pp = Val.create "ltac2:pp" -let val_sort = Val.create "ltac2:sort" -let val_cast = Val.create "ltac2:cast" -let val_inductive = Val.create "ltac2:inductive" -let val_constant = Val.create "ltac2:constant" -let val_constructor = Val.create "ltac2:constructor" -let val_projection = Val.create "ltac2:projection" -let val_univ = Val.create "ltac2:universe" -let val_kont : (Exninfo.iexn -> valexpr Proofview.tactic) Val.typ = - Val.create "ltac2:kont" - -let extract_val (type a) (tag : a Val.typ) (Val.Dyn (tag', v)) : a = -match Val.eq tag tag' with -| None -> assert false -| Some Refl -> v +let coq_core n = KerName.make2 Tac2env.coq_prefix (Label.of_id (Id.of_string_soft n)) module Core = struct @@ -67,108 +46,19 @@ end open Core let v_unit = ValInt 0 -let v_nil = ValInt 0 -let v_cons v vl = ValBlk (0, [|v; vl|]) - -module Value = -struct - -let of_unit () = v_unit - -let to_unit = function -| ValInt 0 -> () -| _ -> assert false - -let of_int n = ValInt n -let to_int = function -| ValInt n -> n -| _ -> assert false - -let of_bool b = if b then ValInt 0 else ValInt 1 - -let to_bool = function -| ValInt 0 -> true -| ValInt 1 -> false -| _ -> assert false - -let of_char n = ValInt (Char.code n) -let to_char = function -| ValInt n -> Char.chr n -| _ -> assert false - -let of_string s = ValStr s -let to_string = function -| ValStr s -> s -| _ -> assert false - -let rec of_list = function -| [] -> v_nil -| x :: l -> v_cons x (of_list l) - -let rec to_list = function -| ValInt 0 -> [] -| ValBlk (0, [|v; vl|]) -> v :: to_list vl -| _ -> assert false - -let of_ext tag c = - ValExt (Val.Dyn (tag, c)) - -let to_ext tag = function -| ValExt e -> extract_val tag e -| _ -> assert false - -let of_constr c = of_ext val_constr c -let to_constr c = to_ext val_constr c - -let of_ident c = of_ext val_ident c -let to_ident c = to_ext val_ident c - -let of_pattern c = of_ext val_pattern c -let to_pattern c = to_ext val_pattern c - -(** FIXME: handle backtrace in Ltac2 exceptions *) -let of_exn c = match fst c with -| LtacError (kn, c) -> ValOpn (kn, c) -| _ -> of_ext val_exn c - -let to_exn c = match c with -| ValOpn (kn, c) -> (LtacError (kn, c), Exninfo.null) -| _ -> to_ext val_exn c - -let of_option = function -| None -> ValInt 0 -| Some c -> ValBlk (0, [|c|]) - -let to_option = function -| ValInt 0 -> None -| ValBlk (0, [|c|]) -> Some c -| _ -> assert false - -let of_pp c = of_ext val_pp c -let to_pp c = to_ext val_pp c - -let of_tuple cl = ValBlk (0, cl) -let to_tuple = function -| ValBlk (0, cl) -> cl -| _ -> assert false - -let of_array = of_tuple -let to_array = to_tuple let of_name c = match c with -| Anonymous -> of_option None -| Name id -> of_option (Some (of_ident id)) +| Anonymous -> Value.of_option None +| Name id -> Value.of_option (Some (Value.of_ident id)) let of_instance sigma u = let u = Univ.Instance.to_array (EConstr.EInstance.kind sigma u) in - of_array (Array.map (fun v -> of_ext val_univ v) u) + Value.of_array (Array.map (fun v -> Value.of_ext Value.val_univ v) u) let of_rec_declaration (nas, ts, cs) = - (of_array (Array.map of_name nas), - of_array (Array.map of_constr ts), - of_array (Array.map of_constr cs)) - -end + (Value.of_array (Array.map of_name nas), + Value.of_array (Array.map Value.of_constr ts), + Value.of_array (Array.map Value.of_constr cs)) let val_valexpr = Val.create "ltac2:valexpr" @@ -395,28 +285,28 @@ let prm_constr_kind : ml_tactic = function Value.of_array (Array.map Value.of_constr args) |]) | Sort s -> - ValBlk (4, [|Value.of_ext val_sort s|]) + ValBlk (4, [|Value.of_ext Value.val_sort s|]) | Cast (c, k, t) -> ValBlk (5, [| Value.of_constr c; - Value.of_ext val_cast k; + Value.of_ext Value.val_cast k; Value.of_constr t; |]) | Prod (na, t, u) -> ValBlk (6, [| - Value.of_name na; + of_name na; Value.of_constr t; Value.of_constr u; |]) | Lambda (na, t, c) -> ValBlk (7, [| - Value.of_name na; + of_name na; Value.of_constr t; Value.of_constr c; |]) | LetIn (na, b, t, c) -> ValBlk (8, [| - Value.of_name na; + of_name na; Value.of_constr b; Value.of_constr t; Value.of_constr c; @@ -428,18 +318,18 @@ let prm_constr_kind : ml_tactic = function |]) | Const (cst, u) -> ValBlk (10, [| - Value.of_ext val_constant cst; - Value.of_instance sigma u; + Value.of_ext Value.val_constant cst; + of_instance sigma u; |]) | Ind (ind, u) -> ValBlk (11, [| - Value.of_ext val_inductive ind; - Value.of_instance sigma u; + Value.of_ext Value.val_inductive ind; + of_instance sigma u; |]) | Construct (cstr, u) -> ValBlk (12, [| - Value.of_ext val_constructor cstr; - Value.of_instance sigma u; + Value.of_ext Value.val_constructor cstr; + of_instance sigma u; |]) | Case (_, c, t, bl) -> ValBlk (13, [| @@ -448,7 +338,7 @@ let prm_constr_kind : ml_tactic = function Value.of_array (Array.map (fun c -> Value.of_constr c) bl); |]) | Fix ((recs, i), def) -> - let (nas, ts, cs) = Value.of_rec_declaration def in + let (nas, ts, cs) = of_rec_declaration def in ValBlk (14, [| Value.of_array (Array.map Value.of_int recs); Value.of_int i; @@ -457,7 +347,7 @@ let prm_constr_kind : ml_tactic = function cs; |]) | CoFix (i, def) -> - let (nas, ts, cs) = Value.of_rec_declaration def in + let (nas, ts, cs) = of_rec_declaration def in ValBlk (15, [| Value.of_int i; nas; @@ -466,7 +356,7 @@ let prm_constr_kind : ml_tactic = function |]) | Proj (p, c) -> ValBlk (16, [| - Value.of_ext val_projection p; + Value.of_ext Value.val_projection p; Value.of_constr c; |]) end @@ -584,7 +474,7 @@ let prm_case : ml_tactic = function Proofview.tclCASE (thaw f) >>= begin function | Proofview.Next (x, k) -> let k = { - clos_env = Id.Map.singleton k_var (Value.of_ext val_kont k); + clos_env = Id.Map.singleton k_var (Value.of_ext Value.val_kont k); clos_var = [Name e_var]; clos_exp = GTacPrm (prm_apply_kont_h, [GTacVar k_var; GTacVar e_var]); } in @@ -595,7 +485,7 @@ let prm_case : ml_tactic = function (** 'a kont -> exn -> 'a *) let prm_apply_kont : ml_tactic = function -| [k; e] -> (Value.to_ext val_kont k) (Value.to_exn e) +| [k; e] -> (Value.to_ext Value.val_kont k) (Value.to_exn e) | _ -> assert false (** int -> int -> (unit -> 'a) -> 'a *) @@ -768,7 +658,7 @@ let interp_constr flags ist (c, _) = Proofview.V82.wrap_exceptions begin fun () -> let ist = to_lvar ist in let (sigma, c) = understand_ltac flags env sigma ist WithoutTypeConstraint c in - let c = Val.Dyn (val_constr, c) in + let c = Val.Dyn (Value.val_constr, c) in Proofview.Unsafe.tclEVARS sigma >>= fun () -> Proofview.tclUNIT c end @@ -791,7 +681,7 @@ let () = define_ml_object Stdarg.wit_open_constr obj let () = - let interp _ id = return (Val.Dyn (val_ident, id)) in + let interp _ id = return (Val.Dyn (Value.val_ident, id)) in let obj = { ml_type = t_ident; ml_interp = interp; @@ -799,7 +689,7 @@ let () = define_ml_object Stdarg.wit_ident obj let () = - let interp _ c = return (Val.Dyn (val_pattern, c)) in + let interp _ c = return (Val.Dyn (Value.val_pattern, c)) in let obj = { ml_type = t_pattern; ml_interp = interp; diff --git a/src/tac2core.mli b/src/tac2core.mli index 41c79b2c65..07ff6cd539 100644 --- a/src/tac2core.mli +++ b/src/tac2core.mli @@ -25,47 +25,3 @@ val t_string : type_constant val t_array : type_constant end - -(** {5 Ltac2 FFI} *) - -(** These functions allow to convert back and forth between OCaml and Ltac2 - data representation. The [to_*] functions raise an anomaly whenever the data - has not expected shape. *) - -module Value : -sig - -val of_unit : unit -> valexpr -val to_unit : valexpr -> unit - -val of_int : int -> valexpr -val to_int : valexpr -> int - -val of_bool : bool -> valexpr -val to_bool : valexpr -> bool - -val of_char : char -> valexpr -val to_char : valexpr -> char - -val of_list : valexpr list -> valexpr -val to_list : valexpr -> valexpr list - -val of_constr : EConstr.t -> valexpr -val to_constr : valexpr -> EConstr.t - -val of_exn : Exninfo.iexn -> valexpr -val to_exn : valexpr -> Exninfo.iexn - -val of_ident : Id.t -> valexpr -val to_ident : valexpr -> Id.t - -val of_array : valexpr array -> valexpr -val to_array : valexpr -> valexpr array - -val of_tuple : valexpr array -> valexpr -val to_tuple : valexpr -> valexpr array - -val of_option : valexpr option -> valexpr -val to_option : valexpr -> valexpr option - -end diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml new file mode 100644 index 0000000000..5f1341ea80 --- /dev/null +++ b/src/tac2ffi.ml @@ -0,0 +1,122 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* t +| _ -> assert false + +let val_constr = val_tag (topwit Stdarg.wit_constr) +let val_ident = val_tag (topwit Stdarg.wit_ident) +let val_pattern = Val.create "ltac2:pattern" +let val_pp = Val.create "ltac2:pp" +let val_sort = Val.create "ltac2:sort" +let val_cast = Val.create "ltac2:cast" +let val_inductive = Val.create "ltac2:inductive" +let val_constant = Val.create "ltac2:constant" +let val_constructor = Val.create "ltac2:constructor" +let val_projection = Val.create "ltac2:projection" +let val_univ = Val.create "ltac2:universe" +let val_kont : (Exninfo.iexn -> valexpr Proofview.tactic) Val.typ = + Val.create "ltac2:kont" + +let extract_val (type a) (tag : a Val.typ) (Val.Dyn (tag', v)) : a = +match Val.eq tag tag' with +| None -> assert false +| Some Refl -> v + +(** Conversion functions *) + +let of_unit () = ValInt 0 + +let to_unit = function +| ValInt 0 -> () +| _ -> assert false + +let of_int n = ValInt n +let to_int = function +| ValInt n -> n +| _ -> assert false + +let of_bool b = if b then ValInt 0 else ValInt 1 + +let to_bool = function +| ValInt 0 -> true +| ValInt 1 -> false +| _ -> assert false + +let of_char n = ValInt (Char.code n) +let to_char = function +| ValInt n -> Char.chr n +| _ -> assert false + +let of_string s = ValStr s +let to_string = function +| ValStr s -> s +| _ -> assert false + +let rec of_list = function +| [] -> ValInt 0 +| x :: l -> ValBlk (0, [| x; of_list l |]) + +let rec to_list = function +| ValInt 0 -> [] +| ValBlk (0, [|v; vl|]) -> v :: to_list vl +| _ -> assert false + +let of_ext tag c = + ValExt (Val.Dyn (tag, c)) + +let to_ext tag = function +| ValExt e -> extract_val tag e +| _ -> assert false + +let of_constr c = of_ext val_constr c +let to_constr c = to_ext val_constr c + +let of_ident c = of_ext val_ident c +let to_ident c = to_ext val_ident c + +let of_pattern c = of_ext val_pattern c +let to_pattern c = to_ext val_pattern c + +(** FIXME: handle backtrace in Ltac2 exceptions *) +let of_exn c = match fst c with +| LtacError (kn, c) -> ValOpn (kn, c) +| _ -> of_ext val_exn c + +let to_exn c = match c with +| ValOpn (kn, c) -> (LtacError (kn, c), Exninfo.null) +| _ -> to_ext val_exn c + +let of_option = function +| None -> ValInt 0 +| Some c -> ValBlk (0, [|c|]) + +let to_option = function +| ValInt 0 -> None +| ValBlk (0, [|c|]) -> Some c +| _ -> assert false + +let of_pp c = of_ext val_pp c +let to_pp c = to_ext val_pp c + +let of_tuple cl = ValBlk (0, cl) +let to_tuple = function +| ValBlk (0, cl) -> cl +| _ -> assert false + +let of_array = of_tuple +let to_array = to_tuple diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli new file mode 100644 index 0000000000..d72d0452a0 --- /dev/null +++ b/src/tac2ffi.mli @@ -0,0 +1,78 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* valexpr +val to_unit : valexpr -> unit + +val of_int : int -> valexpr +val to_int : valexpr -> int + +val of_bool : bool -> valexpr +val to_bool : valexpr -> bool + +val of_char : char -> valexpr +val to_char : valexpr -> char + +val of_string : string -> valexpr +val to_string : valexpr -> string + +val of_list : valexpr list -> valexpr +val to_list : valexpr -> valexpr list + +val of_constr : EConstr.t -> valexpr +val to_constr : valexpr -> EConstr.t + +val of_exn : Exninfo.iexn -> valexpr +val to_exn : valexpr -> Exninfo.iexn + +val of_ident : Id.t -> valexpr +val to_ident : valexpr -> Id.t + +val of_array : valexpr array -> valexpr +val to_array : valexpr -> valexpr array + +val of_tuple : valexpr array -> valexpr +val to_tuple : valexpr -> valexpr array + +val of_option : valexpr option -> valexpr +val to_option : valexpr -> valexpr option + +val of_pattern : Pattern.constr_pattern -> valexpr +val to_pattern : valexpr -> Pattern.constr_pattern + +val of_pp : Pp.t -> valexpr +val to_pp : valexpr -> Pp.t + +val of_ext : 'a Val.typ -> 'a -> valexpr +val to_ext : 'a Val.typ -> valexpr -> 'a + +(** {5 Dynamic tags} *) + +val val_constr : EConstr.t Val.typ +val val_ident : Id.t Val.typ +val val_pattern : Pattern.constr_pattern Val.typ +val val_pp : Pp.t Val.typ +val val_sort : ESorts.t Val.typ +val val_cast : Constr.cast_kind Val.typ +val val_inductive : inductive Val.typ +val val_constant : Constant.t Val.typ +val val_constructor : constructor Val.typ +val val_projection : Projection.t Val.typ +val val_univ : Univ.universe_level Val.typ +val val_kont : (Exninfo.iexn -> valexpr Proofview.tactic) Val.typ diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 89a8d98693..25d83c06fb 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -12,12 +12,14 @@ open Tac2expr open Tac2core open Proofview.Notations +module Value = Tac2ffi + (** Standard tactics sharing their implementation with Ltac1 *) let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } let return x = Proofview.tclUNIT x -let v_unit = Tac2core.Value.of_unit () +let v_unit = Value.of_unit () let lift tac = tac <*> return v_unit -- cgit v1.2.3 From 23f10f3a1a0fd6498cad975b39af5dd3a8559f06 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 28 Jul 2017 18:32:22 +0200 Subject: Parameterizing FFI functions for parameterized types. --- src/tac2core.ml | 49 +++++++++++++++++++++---------------------------- src/tac2ffi.ml | 22 ++++++++++++---------- src/tac2ffi.mli | 12 ++++++------ src/tac2stdlib.ml | 17 ++++++++--------- 4 files changed, 47 insertions(+), 53 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 6d9ede4421..f28c5c5dcf 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -48,17 +48,17 @@ open Core let v_unit = ValInt 0 let of_name c = match c with -| Anonymous -> Value.of_option None -| Name id -> Value.of_option (Some (Value.of_ident id)) +| Anonymous -> Value.of_option Value.of_ident None +| Name id -> Value.of_option Value.of_ident (Some id) let of_instance sigma u = let u = Univ.Instance.to_array (EConstr.EInstance.kind sigma u) in - Value.of_array (Array.map (fun v -> Value.of_ext Value.val_univ v) u) + Value.of_array (fun v -> Value.of_ext Value.val_univ v) u let of_rec_declaration (nas, ts, cs) = - (Value.of_array (Array.map of_name nas), - Value.of_array (Array.map Value.of_constr ts), - Value.of_array (Array.map Value.of_constr cs)) + (Value.of_array of_name nas, + Value.of_array Value.of_constr ts, + Value.of_array Value.of_constr cs) let val_valexpr = Val.create "ltac2:valexpr" @@ -183,11 +183,8 @@ let prm_ident_to_string : ml_tactic = function let prm_ident_of_string : ml_tactic = function | [s] -> let s = Value.to_string s in - let id = - try Value.of_option (Some (Value.of_ident (Id.of_string s))) - with _ -> Value.of_option None - in - return id + let id = try Some (Id.of_string s) with _ -> None in + return (Value.of_option Value.of_ident id) | _ -> assert false (** Int *) @@ -282,7 +279,7 @@ let prm_constr_kind : ml_tactic = function | Evar (evk, args) -> ValBlk (3, [| Value.of_int (Evar.repr evk); - Value.of_array (Array.map Value.of_constr args) + Value.of_array Value.of_constr args; |]) | Sort s -> ValBlk (4, [|Value.of_ext Value.val_sort s|]) @@ -314,7 +311,7 @@ let prm_constr_kind : ml_tactic = function | App (c, cl) -> ValBlk (9, [| Value.of_constr c; - Value.of_array (Array.map Value.of_constr cl) + Value.of_array Value.of_constr cl; |]) | Const (cst, u) -> ValBlk (10, [| @@ -335,12 +332,12 @@ let prm_constr_kind : ml_tactic = function ValBlk (13, [| Value.of_constr c; Value.of_constr t; - Value.of_array (Array.map (fun c -> Value.of_constr c) bl); + Value.of_array Value.of_constr bl; |]) | Fix ((recs, i), def) -> let (nas, ts, cs) = of_rec_declaration def in ValBlk (14, [| - Value.of_array (Array.map Value.of_int recs); + Value.of_array Value.of_int recs; Value.of_int i; nas; ts; @@ -378,7 +375,7 @@ let prm_pattern_matches : ml_tactic = function | Some ans -> let ans = Id.Map.bindings ans in let of_pair (id, c) = Value.of_tuple [| Value.of_ident id; Value.of_constr c |] in - return (Value.of_list (List.map of_pair ans)) + return (Value.of_list of_pair ans) end end | _ -> assert false @@ -393,7 +390,7 @@ let prm_pattern_matches_subterm : ml_tactic = function | IStream.Cons ({ m_sub = (_, sub); m_ctx }, s) -> let ans = Id.Map.bindings sub in let of_pair (id, c) = Value.of_tuple [| Value.of_ident id; Value.of_constr c |] in - let ans = Value.of_tuple [| Value.of_constr m_ctx; Value.of_list (List.map of_pair ans) |] in + let ans = Value.of_tuple [| Value.of_constr m_ctx; Value.of_list of_pair ans |] in Proofview.tclOR (return ans) (fun _ -> of_ans s) in pf_apply begin fun env sigma -> @@ -441,19 +438,16 @@ let prm_once : ml_tactic = function (** (unit -> unit) list -> unit *) let prm_dispatch : ml_tactic = function | [l] -> - let l = Value.to_list l in - let l = List.map (fun f -> Proofview.tclIGNORE (thaw f)) l in + let l = Value.to_list (fun f -> Proofview.tclIGNORE (thaw f)) l in Proofview.tclDISPATCH l >>= fun () -> return v_unit | _ -> assert false (** (unit -> unit) list -> (unit -> unit) -> (unit -> unit) list -> unit *) let prm_extend : ml_tactic = function | [lft; tac; rgt] -> - let lft = Value.to_list lft in - let lft = List.map (fun f -> Proofview.tclIGNORE (thaw f)) lft in + let lft = Value.to_list (fun f -> Proofview.tclIGNORE (thaw f)) lft in let tac = Proofview.tclIGNORE (thaw tac) in - let rgt = Value.to_list rgt in - let rgt = List.map (fun f -> Proofview.tclIGNORE (thaw f)) rgt in + let rgt = Value.to_list (fun f -> Proofview.tclIGNORE (thaw f)) rgt in Proofview.tclEXTEND lft tac rgt >>= fun () -> return v_unit | _ -> assert false @@ -540,18 +534,17 @@ let prm_hyps : ml_tactic = function | [_] -> pf_apply begin fun env _ -> let open Context.Named.Declaration in - let hyps = Environ.named_context env in + let hyps = List.rev (Environ.named_context env) in let map = function | LocalAssum (id, t) -> let t = EConstr.of_constr t in - Value.of_tuple [|Value.of_ident id; Value.of_option None; Value.of_constr t|] + Value.of_tuple [|Value.of_ident id; Value.of_option Value.of_constr None; Value.of_constr t|] | LocalDef (id, c, t) -> let c = EConstr.of_constr c in let t = EConstr.of_constr t in - Value.of_tuple [|Value.of_ident id; Value.of_option (Some (Value.of_constr c)); Value.of_constr t|] + Value.of_tuple [|Value.of_ident id; Value.of_option Value.of_constr (Some c); Value.of_constr t|] in - let hyps = List.rev_map map hyps in - return (Value.of_list hyps) + return (Value.of_list map hyps) end | _ -> assert false diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index 5f1341ea80..74e2b02aeb 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -67,13 +67,13 @@ let to_string = function | ValStr s -> s | _ -> assert false -let rec of_list = function +let rec of_list f = function | [] -> ValInt 0 -| x :: l -> ValBlk (0, [| x; of_list l |]) +| x :: l -> ValBlk (0, [| f x; of_list f l |]) -let rec to_list = function +let rec to_list f = function | ValInt 0 -> [] -| ValBlk (0, [|v; vl|]) -> v :: to_list vl +| ValBlk (0, [|v; vl|]) -> f v :: to_list f vl | _ -> assert false let of_ext tag c = @@ -101,13 +101,13 @@ let to_exn c = match c with | ValOpn (kn, c) -> (LtacError (kn, c), Exninfo.null) | _ -> to_ext val_exn c -let of_option = function +let of_option f = function | None -> ValInt 0 -| Some c -> ValBlk (0, [|c|]) +| Some c -> ValBlk (0, [|f c|]) -let to_option = function +let to_option f = function | ValInt 0 -> None -| ValBlk (0, [|c|]) -> Some c +| ValBlk (0, [|c|]) -> Some (f c) | _ -> assert false let of_pp c = of_ext val_pp c @@ -118,5 +118,7 @@ let to_tuple = function | ValBlk (0, cl) -> cl | _ -> assert false -let of_array = of_tuple -let to_array = to_tuple +let of_array f vl = ValBlk (0, Array.map f vl) +let to_array f = function +| ValBlk (0, vl) -> Array.map f vl +| _ -> assert false diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index d72d0452a0..3f429995ce 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -32,8 +32,8 @@ val to_char : valexpr -> char val of_string : string -> valexpr val to_string : valexpr -> string -val of_list : valexpr list -> valexpr -val to_list : valexpr -> valexpr list +val of_list : ('a -> valexpr) -> 'a list -> valexpr +val to_list : (valexpr -> 'a) -> valexpr -> 'a list val of_constr : EConstr.t -> valexpr val to_constr : valexpr -> EConstr.t @@ -44,14 +44,14 @@ val to_exn : valexpr -> Exninfo.iexn val of_ident : Id.t -> valexpr val to_ident : valexpr -> Id.t -val of_array : valexpr array -> valexpr -val to_array : valexpr -> valexpr array +val of_array : ('a -> valexpr) -> 'a array -> valexpr +val to_array : (valexpr -> 'a) -> valexpr -> 'a array val of_tuple : valexpr array -> valexpr val to_tuple : valexpr -> valexpr array -val of_option : valexpr option -> valexpr -val to_option : valexpr -> valexpr option +val of_option : ('a -> valexpr) -> 'a option -> valexpr +val to_option : (valexpr -> 'a) -> valexpr -> 'a option val of_pattern : Pattern.constr_pattern -> valexpr val to_pattern : valexpr -> Pattern.constr_pattern diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 25d83c06fb..0aeccbd9c6 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -112,45 +112,44 @@ let () = define_prim0 "tac_split" (Tactics.split_with_bindings false [NoBindings let () = define_prim0 "tac_esplit" (Tactics.split_with_bindings true [NoBindings]) let () = define_prim1 "tac_rename" begin fun ids -> - let ids = Value.to_list ids in let map c = match Value.to_tuple c with | [|x; y|] -> (Value.to_ident x, Value.to_ident y) | _ -> assert false in - let ids = List.map map ids in + let ids = Value.to_list map ids in Tactics.rename_hyp ids end let () = define_prim1 "tac_revert" begin fun ids -> - let ids = List.map Value.to_ident (Value.to_list ids) in + let ids = Value.to_list Value.to_ident ids in Tactics.revert ids end let () = define_prim0 "tac_admit" Proofview.give_up let () = define_prim2 "tac_fix" begin fun idopt n -> - let idopt = Option.map Value.to_ident (Value.to_option idopt) in + let idopt = Value.to_option Value.to_ident idopt in let n = Value.to_int n in Tactics.fix idopt n end let () = define_prim1 "tac_cofix" begin fun idopt -> - let idopt = Option.map Value.to_ident (Value.to_option idopt) in + let idopt = Value.to_option Value.to_ident idopt in Tactics.cofix idopt end let () = define_prim1 "tac_clear" begin fun ids -> - let ids = List.map Value.to_ident (Value.to_list ids) in + let ids = Value.to_list Value.to_ident ids in Tactics.clear ids end let () = define_prim1 "tac_keep" begin fun ids -> - let ids = List.map Value.to_ident (Value.to_list ids) in + let ids = Value.to_list Value.to_ident ids in Tactics.keep ids end let () = define_prim1 "tac_clearbody" begin fun ids -> - let ids = List.map Value.to_ident (Value.to_list ids) in + let ids = Value.to_list Value.to_ident ids in Tactics.clear_body ids end @@ -161,7 +160,7 @@ let () = define_prim1 "tac_absurd" begin fun c -> end let () = define_prim1 "tac_subst" begin fun ids -> - let ids = List.map Value.to_ident (Value.to_list ids) in + let ids = Value.to_list Value.to_ident ids in Equality.subst ids end -- cgit v1.2.3 From 0f72b089de52ad7d26d71e56003b140fa5012635 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 28 Jul 2017 17:53:42 +0200 Subject: Exporting more internals from Coq implementation. --- src/tac2core.ml | 9 ++++ src/tac2env.ml | 3 ++ src/tac2env.mli | 3 ++ src/tac2stdlib.ml | 125 +++++++++++++++++++++++++++++++++++++++++++++++----- tests/stuff/ltac2.v | 8 +++- theories/Control.v | 7 +++ theories/Std.v | 51 +++++++++++++++++---- 7 files changed, 185 insertions(+), 21 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index f28c5c5dcf..515cadc525 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -557,6 +557,14 @@ let prm_refine : ml_tactic = function end >>= fun () -> return v_unit | _ -> assert false +let prm_with_holes : ml_tactic = function +| [x; f] -> + Proofview.tclEVARMAP >>= fun sigma0 -> + thaw x >>= fun ans -> + Proofview.tclEVARMAP >>= fun sigma -> + Proofview.Unsafe.tclEVARS sigma0 >>= fun () -> + Tacticals.New.tclWITHHOLES false (interp_app f [ans]) sigma +| _ -> assert false (** Registering *) @@ -615,6 +623,7 @@ let () = Tac2env.define_primitive (pname "goal") prm_goal let () = Tac2env.define_primitive (pname "hyp") prm_hyp let () = Tac2env.define_primitive (pname "hyps") prm_hyps let () = Tac2env.define_primitive (pname "refine") prm_refine +let () = Tac2env.define_primitive (pname "with_holes") prm_with_holes (** ML types *) diff --git a/src/tac2env.ml b/src/tac2env.ml index 2094898ced..a75500eae7 100644 --- a/src/tac2env.ml +++ b/src/tac2env.ml @@ -239,6 +239,9 @@ let interp_ml_object t = MLType.obj t let coq_prefix = MPfile (DirPath.make (List.map Id.of_string ["Init"; "Ltac2"])) +let std_prefix = + MPfile (DirPath.make (List.map Id.of_string ["Std"; "Ltac2"])) + (** Generic arguments *) let wit_ltac2 = Genarg.make0 "ltac2" diff --git a/src/tac2env.mli b/src/tac2env.mli index e26109b691..fea03c4285 100644 --- a/src/tac2env.mli +++ b/src/tac2env.mli @@ -106,6 +106,9 @@ val interp_ml_object : ('a, 'b, 'c) genarg_type -> 'b ml_object val coq_prefix : ModPath.t (** Path where primitive datatypes are defined in Ltac2 plugin. *) +val std_prefix : ModPath.t +(** Path where Ltac-specific datatypes are defined in Ltac2 plugin. *) + (** {5 Generic arguments} *) val wit_ltac2 : (raw_tacexpr, glb_tacexpr, Util.Empty.t) genarg_type diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 0aeccbd9c6..7c7b539113 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open Names open Locus open Misctypes open Tac2expr @@ -14,6 +15,57 @@ open Proofview.Notations module Value = Tac2ffi +let to_pair f g = function +| ValBlk (0, [| x; y |]) -> (f x, g y) +| _ -> assert false + +let to_name c = match Value.to_option Value.to_ident c with +| None -> Anonymous +| Some id -> Name id + +let to_qhyp = function +| ValBlk (0, [| i |]) -> AnonHyp (Value.to_int i) +| ValBlk (1, [| id |]) -> NamedHyp (Value.to_ident id) +| _ -> assert false + +let to_bindings = function +| ValInt 0 -> NoBindings +| ValBlk (0, [| vl |]) -> + ImplicitBindings (Value.to_list Value.to_constr vl) +| ValBlk (1, [| vl |]) -> + ExplicitBindings ((Value.to_list (fun p -> None, to_pair to_qhyp Value.to_constr p) vl)) +| _ -> assert false + +let to_constr_with_bindings = function +| ValBlk (0, [| c; bnd |]) -> (Value.to_constr c, to_bindings bnd) +| _ -> assert false + +let to_int_or_var i = ArgArg (Value.to_int i) + +let to_occurrences f = function +| ValInt 0 -> AllOccurrences +| ValBlk (0, [| vl |]) -> AllOccurrencesBut (Value.to_list f vl) +| ValInt 1 -> NoOccurrences +| ValBlk (1, [| vl |]) -> OnlyOccurrences (Value.to_list f vl) +| _ -> assert false + +let to_hyp_location_flag = function +| ValInt 0 -> InHyp +| ValInt 1 -> InHypTypeOnly +| ValInt 2 -> InHypValueOnly +| _ -> assert false + +let to_clause = function +| ValBlk (0, [| hyps; concl |]) -> + let cast = function + | ValBlk (0, [| hyp; occ; flag |]) -> + ((to_occurrences to_int_or_var occ, Value.to_ident hyp), to_hyp_location_flag flag) + | _ -> assert false + in + let hyps = Value.to_option (fun h -> Value.to_list cast h) hyps in + { onhyps = hyps; concl_occs = to_occurrences to_int_or_var concl; } +| _ -> assert false + (** Standard tactics sharing their implementation with Ltac1 *) let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } @@ -50,6 +102,29 @@ let define_prim2 name tac = in Tac2env.define_primitive (pname name) tac +(** Tactics from Tacexpr *) + +let () = define_prim2 "tac_eelim" begin fun c copt -> + let c = to_constr_with_bindings c in + let copt = Value.to_option to_constr_with_bindings copt in + Tactics.elim true None c copt +end + +let () = define_prim1 "tac_ecase" begin fun c -> + let c = to_constr_with_bindings c in + Tactics.general_case_analysis true None c +end + +let () = define_prim1 "tac_egeneralize" begin fun cl -> + let cast = function + | ValBlk (0, [| c; occs; na |]) -> + ((to_occurrences Value.to_int c, Value.to_constr c), to_name na) + | _ -> assert false + in + let cl = Value.to_list cast cl in + Tactics.new_generalize_gen cl +end + (** Tactics from coretactics *) let () = define_prim0 "tac_reflexivity" Tactics.intros_reflexivity @@ -76,10 +151,26 @@ let () = define_prim1 "tac_cut" begin fun c -> Tactics.cut c end -let () = define_prim0 "tac_left" (Tactics.left_with_bindings false NoBindings) -let () = define_prim0 "tac_eleft" (Tactics.left_with_bindings true NoBindings) -let () = define_prim0 "tac_right" (Tactics.right_with_bindings false NoBindings) -let () = define_prim0 "tac_eright" (Tactics.right_with_bindings true NoBindings) +let () = define_prim1 "tac_left" begin fun bnd -> + let bnd = to_bindings bnd in + Tactics.left_with_bindings false bnd +end +let () = define_prim1 "tac_eleft" begin fun bnd -> + let bnd = to_bindings bnd in + Tactics.left_with_bindings true bnd +end +let () = define_prim1 "tac_right" begin fun bnd -> + let bnd = to_bindings bnd in + Tactics.right_with_bindings false bnd +end +let () = define_prim1 "tac_eright" begin fun bnd -> + let bnd = to_bindings bnd in + Tactics.right_with_bindings true bnd +end + +let () = define_prim1 "tac_introsuntil" begin fun h -> + Tactics.intros_until (to_qhyp h) +end let () = define_prim1 "tac_exactnocheck" begin fun c -> Tactics.exact_no_check (Value.to_constr c) @@ -96,20 +187,32 @@ end let () = define_prim0 "tac_constructor" (Tactics.any_constructor false None) let () = define_prim0 "tac_econstructor" (Tactics.any_constructor true None) -let () = define_prim1 "tac_constructorn" begin fun n -> +let () = define_prim2 "tac_constructorn" begin fun n bnd -> let n = Value.to_int n in - Tactics.constructor_tac false None n NoBindings + let bnd = to_bindings bnd in + Tactics.constructor_tac false None n bnd end -let () = define_prim1 "tac_econstructorn" begin fun n -> +let () = define_prim2 "tac_econstructorn" begin fun n bnd -> let n = Value.to_int n in - Tactics.constructor_tac true None n NoBindings + let bnd = to_bindings bnd in + Tactics.constructor_tac true None n bnd end -let () = define_prim0 "tac_symmetry" (Tactics.intros_symmetry Locusops.onConcl) +let () = define_prim1 "tac_symmetry" begin fun cl -> + let cl = to_clause cl in + Tactics.intros_symmetry cl +end -let () = define_prim0 "tac_split" (Tactics.split_with_bindings false [NoBindings]) -let () = define_prim0 "tac_esplit" (Tactics.split_with_bindings true [NoBindings]) +let () = define_prim1 "tac_split" begin fun bnd -> + let bnd = to_bindings bnd in + Tactics.split_with_bindings false [bnd] +end + +let () = define_prim1 "tac_esplit" begin fun bnd -> + let bnd = to_bindings bnd in + Tactics.split_with_bindings true [bnd] +end let () = define_prim1 "tac_rename" begin fun ids -> let map c = match Value.to_tuple c with diff --git a/tests/stuff/ltac2.v b/tests/stuff/ltac2.v index 4950a20ec4..35546ef6c1 100644 --- a/tests/stuff/ltac2.v +++ b/tests/stuff/ltac2.v @@ -120,7 +120,6 @@ Abort. Goal True. Proof. - let x () := plus (fun () => 0) (fun _ => 1) in match case x with | Val x => @@ -131,6 +130,13 @@ match case x with end. Abort. +Goal (forall n : nat, n = 0 -> False) -> True. +Proof. +refine (fun () => '(fun H => _)). +Std.ecase (hyp @H, Std.ExplicitBindings [Std.NamedHyp @n, '0]). +refine (fun () => 'eq_refl). +Qed. + Ltac2 rec do n tac := match Int.equal n 0 with | true => () | false => tac (); do (Int.sub n 1) tac diff --git a/theories/Control.v b/theories/Control.v index a6d46a89a8..a8b92aced2 100644 --- a/theories/Control.v +++ b/theories/Control.v @@ -54,3 +54,10 @@ Ltac2 @ external hyps : unit -> (ident * constr option * constr) list := "ltac2" (** Refinement *) Ltac2 @ external refine : (unit -> constr) -> unit := "ltac2" "refine". + +(** Evars *) + +Ltac2 @ external with_holes : (unit -> 'a) -> ('a -> 'b) -> 'b := "ltac2" "with_holes". +(** [with_holes x f] evaluates [x], then apply [f] to the result, and fails if + all evars generated by the call to [x] have not been solved when [f] + returns. *) diff --git a/theories/Std.v b/theories/Std.v index 5cc1622ba9..a9eced6cbb 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -8,8 +8,39 @@ Require Import Ltac2.Init. +(** ML-facing types *) + +Ltac2 Type hypothesis := [ AnonHyp (int) | NamedHyp (ident) ]. + +Ltac2 Type bindings := [ +| NoBindings +| ImplicitBindings (constr list) +| ExplicitBindings ((hypothesis * constr) list) +]. + +Ltac2 Type constr_with_bindings := constr * bindings. + +Ltac2 Type occurrences := [ +| AllOccurrences +| AllOccurrencesBut (int list) +| NoOccurrences +| OnlyOccurrences (int list) +]. + +Ltac2 Type hyp_location_flag := [ InHyp | InHypTypeOnly | InHypValueOnly ]. + +Ltac2 Type clause := { + on_hyps : (ident * occurrences * hyp_location_flag) list option; + on_concl : occurrences; +}. + (** Standard, built-in tactics. See Ltac1 for documentation. *) +Ltac2 @ external eelim : constr_with_bindings -> constr_with_bindings option -> unit := "ltac2" "tac_eelim". +Ltac2 @ external ecase : constr_with_bindings -> unit := "ltac2" "tac_ecase". + +Ltac2 @ external egeneralize : (constr * occurrences * ident option) list -> unit := "ltac2" "tac_egeneralize". + Ltac2 @ external reflexivity : unit -> unit := "ltac2" "tac_reflexivity". Ltac2 @ external assumption : unit -> unit := "ltac2" "tac_assumption". @@ -20,20 +51,22 @@ Ltac2 @ external etransitivity : unit -> unit := "ltac2" "tac_etransitivity". Ltac2 @ external cut : constr -> unit := "ltac2" "tac_cut". -Ltac2 @ external left : unit -> unit := "ltac2" "tac_left". -Ltac2 @ external eleft : unit -> unit := "ltac2" "tac_eleft". -Ltac2 @ external right : unit -> unit := "ltac2" "tac_right". -Ltac2 @ external eright : unit -> unit := "ltac2" "tac_eright". +Ltac2 @ external left : bindings -> unit := "ltac2" "tac_left". +Ltac2 @ external eleft : bindings -> unit := "ltac2" "tac_eleft". +Ltac2 @ external right : bindings -> unit := "ltac2" "tac_right". +Ltac2 @ external eright : bindings -> unit := "ltac2" "tac_eright". Ltac2 @ external constructor : unit -> unit := "ltac2" "tac_constructor". Ltac2 @ external econstructor : unit -> unit := "ltac2" "tac_econstructor". -Ltac2 @ external split : unit -> unit := "ltac2" "tac_split". -Ltac2 @ external esplit : unit -> unit := "ltac2" "tac_esplit". +Ltac2 @ external split : bindings -> unit := "ltac2" "tac_split". +Ltac2 @ external esplit : bindings -> unit := "ltac2" "tac_esplit". + +Ltac2 @ external constructor_n : int -> bindings -> unit := "ltac2" "tac_constructorn". +Ltac2 @ external econstructor_n : int -> bindings -> unit := "ltac2" "tac_econstructorn". -Ltac2 @ external constructor_n : int -> unit := "ltac2" "tac_constructorn". -Ltac2 @ external econstructor_n : int -> unit := "ltac2" "tac_econstructorn". +Ltac2 @ external intros_until : hypothesis -> unit := "ltac2" "tac_introsuntil". -Ltac2 @ external symmetry : unit -> unit := "ltac2" "tac_symmetry". +Ltac2 @ external symmetry : clause -> unit := "ltac2" "tac_symmetry". Ltac2 @ external rename : (ident * ident) list -> unit := "ltac2" "tac_rename". -- cgit v1.2.3 From f9e7c43b5884f5231f14ec7b008b1eb660026a0e Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 31 Jul 2017 22:07:55 +0200 Subject: Adding new scopes for standard Ltac structures. --- _CoqProject | 2 ++ src/g_ltac2.ml4 | 63 +++++++++++++++++++++++++++++++++++++++++++++---- src/ltac2_plugin.mlpack | 1 + src/tac2core.ml | 27 ++++++++++++++++++++- src/tac2core.mli | 13 ++++++++-- src/tac2entries.ml | 3 +++ src/tac2entries.mli | 5 ++++ src/tac2quote.ml | 63 +++++++++++++++++++++++++++++++++++++++++++++++++ src/tac2quote.mli | 32 +++++++++++++++++++++++++ tests/example2.v | 15 ++++++++++++ 10 files changed, 216 insertions(+), 8 deletions(-) create mode 100644 src/tac2quote.ml create mode 100644 src/tac2quote.mli create mode 100644 tests/example2.v diff --git a/_CoqProject b/_CoqProject index 6c9393628d..ab73af1295 100644 --- a/_CoqProject +++ b/_CoqProject @@ -17,6 +17,8 @@ src/tac2ffi.ml src/tac2ffi.mli src/tac2core.ml src/tac2core.mli +src/tac2quote.ml +src/tac2quote.mli src/tac2stdlib.ml src/tac2stdlib.mli src/g_ltac2.ml4 diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 21612f9a25..4a2f615df9 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -9,12 +9,29 @@ open Pp open Util open Names +open Tok open Pcoq open Constrexpr open Misctypes open Tac2expr open Ltac_plugin +let err () = raise Stream.Failure + +(* idem for (x:=t) and (1:=t) *) +let test_lpar_idnum_coloneq = + Gram.Entry.of_parser "test_lpar_idnum_coloneq" + (fun strm -> + match stream_nth 0 strm with + | KEYWORD "(" -> + (match stream_nth 1 strm with + | IDENT _ | INT _ -> + (match stream_nth 2 strm with + | KEYWORD ":=" -> () + | _ -> err ()) + | _ -> err ()) + | _ -> err ()) + let tac2expr = Tac2entries.Pltac.tac2expr let tac2type = Gram.entry_create "tactic:tac2type" let tac2def_val = Gram.entry_create "tactic:tac2def_val" @@ -24,11 +41,12 @@ let tac2def_syn = Gram.entry_create "tactic:tac2def_syn" let tac2mode = Gram.entry_create "vernac:ltac2_command" let inj_wit wit loc x = CTacExt (loc, Genarg.in_gen (Genarg.rawwit wit) x) -let inj_constr loc c = inj_wit Stdarg.wit_constr loc c let inj_open_constr loc c = inj_wit Stdarg.wit_open_constr loc c -let inj_ident loc c = inj_wit Stdarg.wit_ident loc c let inj_pattern loc c = inj_wit Tac2env.wit_pattern loc c +let mk_constr ~loc kn args = + CTacApp (loc, CTacCst (loc, AbsKn (Other kn)), args) + let pattern_of_qualid loc id = if Tac2env.is_constructor (snd id) then CPatRef (loc, RelId id, []) else @@ -108,11 +126,11 @@ GEXTEND Gram | s = Prim.string -> CTacAtm (Loc.tag ~loc:!@loc (AtmStr s)) | id = Prim.qualid -> if Tac2env.is_constructor (snd id) then CTacCst (!@loc, RelId id) else CTacRef (RelId id) - | "@"; id = Prim.ident -> inj_ident !@loc id + | "@"; id = Prim.ident -> Tac2quote.of_ident ~loc:!@loc id | "'"; c = Constr.constr -> inj_open_constr !@loc c - | IDENT "constr"; ":"; "("; c = Constr.lconstr; ")" -> inj_constr !@loc c + | IDENT "constr"; ":"; "("; c = Constr.lconstr; ")" -> Tac2quote.of_constr ~loc:!@loc c | IDENT "open_constr"; ":"; "("; c = Constr.lconstr; ")" -> inj_open_constr !@loc c - | IDENT "ident"; ":"; "("; c = Prim.ident; ")" -> inj_ident !@loc c + | IDENT "ident"; ":"; "("; c = Prim.ident; ")" -> Tac2quote.of_ident ~loc:!@loc c | IDENT "pattern"; ":"; "("; c = Constr.lconstr_pattern; ")" -> inj_pattern !@loc c ] ] ; @@ -256,6 +274,41 @@ GEXTEND Gram ; END +(** Quotation scopes used by notations *) + +open Tac2entries.Pltac + +GEXTEND Gram + GLOBAL: q_ident q_bindings; + q_ident: + [ [ id = Prim.ident -> Tac2quote.of_ident ~loc:!@loc id + | "$"; id = Prim.ident -> Tac2quote.of_variable ~loc:!@loc id + ] ] + ; + simple_binding: + [ [ "("; id = Prim.ident; ":="; c = Constr.lconstr; ")" -> + Loc.tag ~loc:!@loc (NamedHyp id, Tac2quote.of_constr ~loc:!@loc c) + | "("; n = Prim.natural; ":="; c = Constr.lconstr; ")" -> + Loc.tag ~loc:!@loc (AnonHyp n, Tac2quote.of_constr ~loc:!@loc c) + ] ] + ; + bindings: + [ [ test_lpar_idnum_coloneq; bl = LIST1 simple_binding -> + Tac2quote.of_bindings ~loc:!@loc (ExplicitBindings bl) + | bl = LIST1 Constr.constr -> + let bl = List.map (fun c -> Tac2quote.of_constr ~loc:!@loc c) bl in + Tac2quote.of_bindings ~loc:!@loc (Misctypes.ImplicitBindings bl) + ] ] + ; + q_bindings: + [ [ "with"; bl = bindings -> bl + | -> mk_constr ~loc:!@loc Tac2core.Core.c_no_bindings [] + ] ] + ; +END + +(** Extension of constr syntax *) + GEXTEND Gram Pcoq.Constr.operconstr: LEVEL "0" [ [ IDENT "ltac2"; ":"; "("; tac = tac2expr; ")" -> diff --git a/src/ltac2_plugin.mlpack b/src/ltac2_plugin.mlpack index 1d7b655dce..8d2d7dc0f4 100644 --- a/src/ltac2_plugin.mlpack +++ b/src/ltac2_plugin.mlpack @@ -5,5 +5,6 @@ Tac2interp Tac2entries Tac2ffi Tac2core +Tac2quote Tac2stdlib G_ltac2 diff --git a/src/tac2core.ml b/src/tac2core.ml index 515cadc525..111ef1c8eb 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -21,6 +21,7 @@ open Proofview.Notations module Value = Tac2ffi let coq_core n = KerName.make2 Tac2env.coq_prefix (Label.of_id (Id.of_string_soft n)) +let std_core n = KerName.make2 Tac2env.std_prefix (Label.of_id (Id.of_string_soft n)) module Core = struct @@ -41,6 +42,15 @@ let c_cons = coq_core "::" let c_none = coq_core "None" let c_some = coq_core "Some" +let t_bindings = std_core "bindings" +let c_no_bindings = std_core "NoBindings" +let c_implicit_bindings = std_core "ImplicitBindings" +let c_explicit_bindings = std_core "ExplicitBindings" + +let t_qhyp = std_core "hypothesis" +let c_named_hyp = std_core "NamedHyp" +let c_anon_hyp = std_core "AnonHyp" + end open Core @@ -835,5 +845,20 @@ let () = add_scope "tactic" begin function | _ -> scope_fail () end -let () = add_generic_scope "ident" Pcoq.Prim.ident Stdarg.wit_ident +let () = add_scope "ident" begin function +| [] -> + let scope = Extend.Aentry Tac2entries.Pltac.q_ident in + let act tac = rthunk tac in + Tac2entries.ScopeRule (scope, act) +| _ -> scope_fail () +end + +let () = add_scope "bindings" begin function +| [] -> + let scope = Extend.Aentry Tac2entries.Pltac.q_bindings in + let act tac = rthunk tac in + Tac2entries.ScopeRule (scope, act) +| _ -> scope_fail () +end + let () = add_generic_scope "constr" Pcoq.Constr.constr Stdarg.wit_constr diff --git a/src/tac2core.mli b/src/tac2core.mli index 07ff6cd539..118b7aaa42 100644 --- a/src/tac2core.mli +++ b/src/tac2core.mli @@ -16,12 +16,21 @@ module Core : sig val t_list : type_constant -val c_nil : ltac_constant -val c_cons : ltac_constant +val c_nil : ltac_constructor +val c_cons : ltac_constructor val t_int : type_constant val t_option : type_constant val t_string : type_constant val t_array : type_constant +val t_bindings : type_constant +val c_no_bindings : ltac_constructor +val c_implicit_bindings : ltac_constant +val c_explicit_bindings : ltac_constant + +val t_qhyp : type_constant +val c_anon_hyp : ltac_constructor +val c_named_hyp : ltac_constructor + end diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 7bc4c75510..d293a87975 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -23,6 +23,9 @@ open Vernacexpr module Pltac = struct let tac2expr = Pcoq.Gram.entry_create "tactic:tac2expr" + +let q_ident = Pcoq.Gram.entry_create "tactic:q_ident" +let q_bindings = Pcoq.Gram.entry_create "tactic:q_bindings" end (** Tactic definition *) diff --git a/src/tac2entries.mli b/src/tac2entries.mli index 71e8150057..4d5a234daf 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -54,4 +54,9 @@ val call : default:bool -> raw_tacexpr -> unit module Pltac : sig val tac2expr : raw_tacexpr Pcoq.Gram.entry + +(** Quoted entries. They directly return an Ltac2 expression *) + +val q_ident : raw_tacexpr Pcoq.Gram.entry +val q_bindings : raw_tacexpr Pcoq.Gram.entry end diff --git a/src/tac2quote.ml b/src/tac2quote.ml new file mode 100644 index 0000000000..2d9521d30c --- /dev/null +++ b/src/tac2quote.ml @@ -0,0 +1,63 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* constructor Core.c_nil [] +| e :: l -> + constructor ?loc Core.c_cons [e; of_list ?loc l] + +let of_qhyp ?loc = function +| AnonHyp n -> constructor Core.c_anon_hyp [of_int ?loc n] +| NamedHyp id -> constructor Core.c_named_hyp [of_ident ?loc id] + +let of_bindings ?loc = function +| NoBindings -> + constructor ?loc Core.c_no_bindings [] +| ImplicitBindings tl -> + constructor ?loc Core.c_implicit_bindings [of_list ?loc tl] +| ExplicitBindings tl -> + let tl = List.map (fun (loc, (qhyp, e)) -> of_pair ?loc (of_qhyp ?loc qhyp, e)) tl in + constructor ?loc Core.c_explicit_bindings [of_list ?loc tl] diff --git a/src/tac2quote.mli b/src/tac2quote.mli new file mode 100644 index 0000000000..ba6a878d50 --- /dev/null +++ b/src/tac2quote.mli @@ -0,0 +1,32 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* ltac_constructor -> raw_tacexpr list -> raw_tacexpr + +val of_int : ?loc:Loc.t -> int -> raw_tacexpr + +val of_pair : ?loc:Loc.t -> raw_tacexpr * raw_tacexpr -> raw_tacexpr + +val of_variable : ?loc:Loc.t -> Id.t -> raw_tacexpr + +val of_ident : ?loc:Loc.t -> Id.t -> raw_tacexpr + +val of_constr : ?loc:Loc.t -> Constrexpr.constr_expr -> raw_tacexpr + +val of_list : ?loc:Loc.t -> raw_tacexpr list -> raw_tacexpr + +val of_bindings : ?loc:Loc.t -> raw_tacexpr bindings -> raw_tacexpr diff --git a/tests/example2.v b/tests/example2.v new file mode 100644 index 0000000000..14a6b68e18 --- /dev/null +++ b/tests/example2.v @@ -0,0 +1,15 @@ +Require Import Ltac2.Ltac2. + +Ltac2 Notation "split" bnd(bindings) := Std.split (bnd ()). + +Goal exists n, n = 0. +Proof. +split with (x := 0). +Std.reflexivity (). +Qed. + +Goal exists n, n = 0. +Proof. +split with 0. +split. +Qed. -- cgit v1.2.3 From 9eb7f16fc890a1bf3a1332332ed349513905ed66 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 1 Aug 2017 00:35:09 +0200 Subject: Do not thunk notations preemptively. One has to rely on the 'thunk' token to produce such thunks. --- src/tac2core.ml | 22 +++++++++++++++------- tests/example2.v | 2 +- tests/stuff/ltac2.v | 2 +- 3 files changed, 17 insertions(+), 9 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 111ef1c8eb..e0a65dde2d 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -753,7 +753,7 @@ let add_generic_scope s entry arg = let parse = function | [] -> let scope = Extend.Aentry entry in - let act x = rthunk (CTacExt (dummy_loc, in_gen (rawwit arg) x)) in + let act x = CTacExt (dummy_loc, in_gen (rawwit arg) x) in Tac2entries.ScopeRule (scope, act) | _ -> scope_fail () in @@ -818,7 +818,7 @@ end let () = add_scope "self" begin function | [] -> let scope = Extend.Aself in - let act tac = rthunk tac in + let act tac = tac in Tac2entries.ScopeRule (scope, act) | _ -> scope_fail () end @@ -826,7 +826,7 @@ end let () = add_scope "next" begin function | [] -> let scope = Extend.Anext in - let act tac = rthunk tac in + let act tac = tac in Tac2entries.ScopeRule (scope, act) | _ -> scope_fail () end @@ -835,12 +835,12 @@ let () = add_scope "tactic" begin function | [] -> (** Default to level 5 parsing *) let scope = Extend.Aentryl (Tac2entries.Pltac.tac2expr, 5) in - let act tac = rthunk tac in + let act tac = tac in Tac2entries.ScopeRule (scope, act) | [SexprInt (loc, n)] -> let () = if n < 0 || n > 5 then scope_fail () in let scope = Extend.Aentryl (Tac2entries.Pltac.tac2expr, n) in - let act tac = rthunk tac in + let act tac = tac in Tac2entries.ScopeRule (scope, act) | _ -> scope_fail () end @@ -848,7 +848,7 @@ end let () = add_scope "ident" begin function | [] -> let scope = Extend.Aentry Tac2entries.Pltac.q_ident in - let act tac = rthunk tac in + let act tac = tac in Tac2entries.ScopeRule (scope, act) | _ -> scope_fail () end @@ -856,7 +856,15 @@ end let () = add_scope "bindings" begin function | [] -> let scope = Extend.Aentry Tac2entries.Pltac.q_bindings in - let act tac = rthunk tac in + let act tac = tac in + Tac2entries.ScopeRule (scope, act) +| _ -> scope_fail () +end + +let () = add_scope "thunk" begin function +| [tok] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let act e = rthunk (act e) in Tac2entries.ScopeRule (scope, act) | _ -> scope_fail () end diff --git a/tests/example2.v b/tests/example2.v index 14a6b68e18..5efbf90b34 100644 --- a/tests/example2.v +++ b/tests/example2.v @@ -1,6 +1,6 @@ Require Import Ltac2.Ltac2. -Ltac2 Notation "split" bnd(bindings) := Std.split (bnd ()). +Ltac2 Notation "split" bnd(bindings) := Std.split bnd. Goal exists n, n = 0. Proof. diff --git a/tests/stuff/ltac2.v b/tests/stuff/ltac2.v index 35546ef6c1..ece6fca06a 100644 --- a/tests/stuff/ltac2.v +++ b/tests/stuff/ltac2.v @@ -111,7 +111,7 @@ Proof. Fail zero (Bar "lol"). Abort. -Ltac2 Notation "refine!" c(constr) := refine c. +Ltac2 Notation "refine!" c(thunk(constr)) := refine c. Goal True. Proof. -- cgit v1.2.3 From 6fdec59bbd3fc67ff3b0c48193201c1739aa7f70 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 1 Aug 2017 00:44:51 +0200 Subject: Adding documentation from the CEP. --- doc/ltac2.md | 642 +++++++++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 642 insertions(+) create mode 100644 doc/ltac2.md diff --git a/doc/ltac2.md b/doc/ltac2.md new file mode 100644 index 0000000000..38a56d3d6b --- /dev/null +++ b/doc/ltac2.md @@ -0,0 +1,642 @@ +# Summary + +The Ltac tactic language is probably one of the ingredients of the success of +Coq, yet it is at the same time its Achilles' heel. Indeed, Ltac: + +- has nothing like intended semantics +- is very non-uniform due to organic growth +- lacks expressivity and requires programming-by-hacking +- is slow +- is error-prone and fragile +- has an intricate implementation + +This has a lot of terrible consequences, most notably the fact that it is never +clear whether some observed behaviour is a bug or a proper one. + +Following the need of users that start developing huge projects relying +critically on Ltac, we believe that we should offer a proper modern language +that features at least the following: + +- at least informal, predictible semantics +- a typing system +- standard programming facilities (i.e. datatypes) + +This document describes the implementation of such a language. The +implementation of Ltac as of Coq 8.7 will be referred to as Ltac1. + +# General design + +There are various alternatives to Ltac1, such that Mtac or Rtac for instance. +While those alternatives can be quite distinct from Ltac1, we designed +Ltac2 to be closest as reasonably possible to Ltac1, while fixing the +aforementioned defects. + +In particular, Ltac2 is: +- a member of the ML family of languages, i.e. + * a call-by-value functional language + * with effects + * together with Hindley-Milner type system +- a language featuring meta-programming facilities for the manipulation of + Coq-side terms +- a language featuring notation facilities to help writing palatable scripts + +We describe more in details each point in the remainder of this document. + +# ML component + +The call-by-value functional language fragment is easy to implement. + +## Type Syntax + +At the level of terms, we simply elaborate on Ltac1 syntax, which is quite +close to e.g. the one of OCaml. Types follow the simply-typed syntax of OCaml. + +``` +TYPE := +| "(" TYPE₀ "," ... "," TYPEₙ ")" TYPECONST +| "(" TYPE₀ "*" ... "*" TYPEₙ ")" +| TYPE₁ "->" TYPE₂ +| TYPEVAR + +TYPECONST := ( MODPATH "." )* LIDENT + +TYPEVAR := "'" LIDENT + +TYPEPARAMS := "(" TYPEVAR₀ "," ... "," TYPEVARₙ ")" +``` + +The set of base types can be extended thanks to the usual ML type +declarations such as algebraic datatypes and records. + +Built-in types include: +- `int`, machine integers (size not specified, in practice inherited from OCaml) +- `string`, mutable strings +- `'a array`, mutable arrays +- `exn`, exceptions +- `constr`, kernel-side terms +- `pattern`, term patterns +- `ident`, well-formed identifiers + +## Type declarations + +One can define new types by the following commands. + +``` +VERNAC ::= +| "Ltac2" "Type" TYPEPARAMS LIDENT +| "Ltac2" "Type" RECFLAG TYPEPARAMS LIDENT ":=" TYPEDEF + +RECFLAG := ( "rec" ) +``` + +The first command defines an abstract type. It has no use for the end user and +is dedicated to types representing data coming from the OCaml world. + +The second command defines a type with a manifest. There are four possible +kinds of such definitions: alias, variant, record and open variant types. + +``` +TYPEDEF := +| TYPE +| "[" CONSTRUCTORDEF₀ "|" ... "|" CONSTRUCTORDEFₙ "]" +| "{" FIELDDEF₀ ";" ... ";" FIELDDEFₙ "}" +| "[" "..." "]" + +CONSTRUCTORDEF := +| IDENT ( "(" TYPE₀ "," ... "," TYPE₀ ")" ) + +FIELDDEF := +| MUTFLAG IDENT ":" TYPE + +MUTFLAG := ( "mutable" ) +``` + +Aliases are just a name for a given type expression and are transparently +unfoldable to it. They cannot be recursive. + +Variants are sum types defined by constructors and eliminated by +pattern-matching. They can be recursive, but the `RECFLAG` must be explicitly +set. Pattern-maching must be exhaustive. + +Records are product types with named fields and eliminated by projection. +Likewise they can be recursive if the `RECFLAG` is set. + +Open variants are a special kind of variant types whose constructors are not +statically defined, but can instead by extended dynamically. A typical example +is the standard `exn` type. Pattern-matching must always include a catch-all +clause. They can be extended by the following command. + +VERNAC ::= +| "Ltac2" "Type" TYPEPARAMS QUALID ":=" "[" CONSTRUCTORDEF "]" + +## Term Syntax + +The syntax of the functional fragment is very close to the one of Ltac1, except +that it adds a true pattern-matching feature. + +``` +VAR := LIDENT + +QUALID := ( MODPATH "." )* LIDENT + +CONSTRUCTOR := UIDENT + +TERM := +| QUALID +| CONSTRUCTOR TERM₀ ... TERMₙ +| TERM TERM₀ ... TERMₙ +| "fun" VAR "=>" TERM +| "let" VAR ":=" TERM "in" TERM +| "let" "rec" VAR ":=" TERM "in" TERM +| "match" TERM "with" BRANCH* "end" +| INT +| STRING +| "[|" TERM₀ ";" ... ";" TERMₙ "|]" +| "(" TERM₀ "," ... "," TERMₙ ")" +| "{" FIELD+ "}" +| TERM "." "(" QUALID ")" +| TERM₁ "." "(" QUALID ")" ":=" TERM₂ +| "["; TERM₀ ";" ... ";" TERMₙ "]" +| TERM₁ "::" TERM₂ +| ... + +BRANCH := +| PATTERN "=>" TERM + +PATTERN := +| VAR +| "_" +| "(" PATTERN₀ "," ... "," PATTERNₙ ")" +| CONSTRUCTOR PATTERN₀ ... PATTERNₙ +| "[" "]" +| PATTERN₁ "::" PATTERN₂ + +FIELD := +| QUALID ":=" TERM + +``` + +In practice, there is some additional syntactic sugar that allows e.g. to +bind a variable and match on it at the same time, in the usual ML style. + +There is a dedicated syntax for list and array litterals. + +Limitations: for now, deep pattern matching is not implemented yet. + +## Reduction + +We use the usual ML call-by-value reduction, with an otherwise unspecified +evaluation order. + +Note that this is already a departure from Ltac1 which uses heuristic to +decide when evaluating an expression, e.g. the following do not evaluate the +same way. + +``` +foo (idtac; let x := 0 in bar) + +foo (let x := 0 in bar) +``` + +Instead of relying on the `idtac` hack, we would now require an explicit thunk +not to compute the argument, and `foo` would have e.g. type +`(unit -> unit) -> unit`. + +``` +foo (fun () -> let x := 0 in bar) +``` + +## Typing + +Typing is strict and follows Hindley-Milner system. We will not implement the +current hackish subtyping semantics, and one will have to resort to conversion +functions. See notations though to make things more palatable. + +In this setting, all usual argument-free tactics have type `unit -> unit`, but +one can return as well a value of type `τ` thanks to terms of type `unit -> τ`, +or take additional arguments. + +## Effects + +Regarding effects, nothing involved here, except that instead of using the +standard IO monad as the ambient effectful world, Ltac2 is going to use the +tactic monad. + +Note that the order of evaluation of application is *not* specified and is +implementation-dependent, as in OCaml. + +We recall that the `Proofview.tactic` monad is essentially a IO monad together +with backtracking state representing the proof state. + +Intuitively a thunk of type `unit -> 'a` can do the following: +- It can perform non-backtracking IO like printing and setting mutable variables +- It can fail in a non-recoverable way +- It can use first-class backtrack. The proper way to figure that is that we + morally have the following isomorphism: + `(unit -> 'a) ~ unit -> ('a + (exn -> 'a))` i.e. thunks can produce a list + of results waiting for failure exceptions. +- It can access a backtracking proof state, made out amongst other things of + the current evar assignation and the list of goals under focus. + +### Standard IO + +The Ltac2 language features non-backtracking IO, notably mutable data and +printing operations. + +Mutable fields of records can be modified using the set syntax + +### Fatal errors + +The Ltac2 language provides non-backtracking exceptions through the +following primitive in module `Control`. + +``` +val throw : exn -> 'a +``` + +Contrarily to backtracking exceptions from the next section, this kind of error +is never caught by backtracking primitives, that is, throwing an exception +destroys the stack. This is materialized by the following equation, where `E` +is an evaluation context. + +``` +E[throw e] ≡ throw e +``` + +There is currently no way to catch such an exception and it is a design choice. +There might be at some future point a way to catch it in a brutal way, +destroying all backtrack and return values. + +### Backtrack + +In Ltac2, we have the following backtracking primitives, defined in the +`Control` module. + +``` +Ltac2 Type 'a result := [ Val ('a) | Err (exn) ]. + +val zero : exn -> 'a +val plus : (unit -> 'a) -> (exn -> 'a) -> 'a +val case : (unit -> 'a) -> ('a * (exn -> 'a)) result +``` + +If one sees thunks as lazy lists, then `zero` is the empty list and `plus` is +list concatenation, while `case` is pattern-matching. + +The backtracking is first-class, i.e. one can write +`plus "x" (fun () -> "y") : string` producing a backtracking string. + +These operations are expected to satisfy a few equations, most notably that they +form a monoid compatible with sequentialization. + +``` +plus t zero ≡ t () +plus (fun () -> zero e) f ≡ f e +plus (plus t f) g ≡ plus t (fun e -> plus (f e) g) + +case (fun () -> zero e) ≡ Err e +case (fun () -> plus (fun () -> t) f) ≡ Val t f + +let x := zero e in u ≡ fail e +let x := plus t f in u ≡ plus (fun () -> let x := t in u) (fun e -> let x := f e in u) + +(t, u, f, g, e values) +``` + +### Goals + +A goal is given by the data of its conclusion and hypotheses, i.e. it can be +represented as `[Γ ⊢ A]`. + +The tactic monad naturally operates over the whole proofview, which may +represent several goals, including none. Thus, there is no such thing as +*the current goal*. Goals are naturally ordered, though. + +It is natural to do the same in Ltac2, but we must provide a way to get access +to a given goal. This is the role of the `enter` primitive, that applies a +tactic to each currently focussed goal in turn. + +``` +val enter : (unit -> unit) -> unit +``` + +It is guaranteed that when evaluating `enter f`, `f` is called with exactly one +goal under focus. Note that `f` may be called several times, or never, depending +on the number of goals under focus before the call to `enter`. + +A more expressive primitive allows to retrieve the data returned by each tactic +and store it in a list. + +``` +val enter_val : (unit -> 'a) -> 'a list +``` + +Accessing the goal data is then implicit in the Ltac2 primitives, and may fail +if the invariants are not respected. The two essential functions for observing +goals are given below. + +``` +val hyp : ident -> constr +val goal : unit -> constr +``` + +The two above functions fail if there is not exactly one goal under focus. +In addition, `hyp` may also fail if there is no hypothesis with the +corresponding name. + +# Meta-programming + +## Overview + +One of the horrendous implementation issues of Ltac is the fact that it is +never clear whether an object refers to the object world or the meta-world. +This is an incredible source of slowness, as the interpretation must be +aware of bound variables and must use heuristics to decide whether a variable +is a proper one or referring to something in the Ltac context. + +Likewise, in Ltac1, constr parsing is implicit, so that `foo 0` is +not `foo` applied to the Ltac integer expression `0` (Ltac does have a +non-first-class notion of integers), but rather the Coq term `Datatypes.O`. + +We should stop doing that! We need to mark when quoting and unquoting, although +we need to do that in a short and elegant way so as not to be too cumbersome +to the user. + +## Syntax example + +Here is a suggestive example of a reasonable syntax. + +``` +let var := "H" in (* a string *) +let c := << fun $var$ => 0 >> (* the Coq term "fun H => 0" *) +let c' := << let x := $c$ in nat >> (* the Coq term "let x := fun H => 0 in nat" *) +... +``` + +## Term quotation + +### Syntax + +It is better to define primitively the quoting syntax to build terms, as this +is more robust to changes. + +``` +t, u ::= ... | << constr >> +``` + +The `constr` datatype have the same syntax as the usual Coq +terms, except that it also allows antiquotations of the form `$t$` whose type +is statically inferred from the position, e.g. + +``` +<< let $t$ := $u$ >> (** [t] is an ident, [u] is a constr *) +``` + +As the term syntax implicitly allows to inject other classes without marking, +antiquotations can refer explicitly to which class they belong to overcome this +limitation. + +``` +<< $ident:t$ >> (** [t] is an ident, and the corresponding constr is [GVar t] *) +<< $ref:t$ >> (** [t] is a reference, and the corresponding constr is [GRef t] *) +``` + +### Semantics + +Interpretation of a quoted constr is done in two phases, internalization and +evaluation. +- During internalization, variables are resolved and antiquotations are + type-checked as Ltac2 terms, effectively producing a `glob_constr` in Coq + implementation terminology, potentially ill-typed as a Coq term. +- During evaluation, a quoted term is fully evaluated to a kernel term, and is + in particular type-checked in the current environment. + +Internalization is part of the static semantics, i.e. it is done at typing +time, while evaluation is part of the dynamic semantics, i.e. it is done when +a term gets effectively computed. + +#### Static semantics + +The typing rule of a quoted constr is given below, where the `eᵢ` refer to +antiquoted terms. + +``` + Γ ⊢ e₁ : unit Γ ⊢ eₙ : unit +==================================== + Γ ⊢ << c{$e₁$, ..., $eₙ$} >> : constr +``` + +Note that the **static** environment of typing of antiquotations is **not** +expanded by the binders from the term. Namely, it means that the following +expression will **not** type-check. +``` +<< fun x : nat => $exact x$ >> +``` + +There is a simple reason for that, which is that the following expression would +not make sense in general. +``` +<< fun x : nat => $clear x; exact x$ >> +``` + +Rather, the tactic writer has to resort to the **dynamic** environment, and must +write instead something that amounts to the following. +``` +<< fun x : nat => $exact (hyp "x")$ >> +``` + +Obviously, we need to provide syntactic sugar to make this tractable. See the +corresponding section for more details. + +#### Dynamic semantics + +Evaluation of a quoted term is described below. +- The quoted term is evaluated by the pretyper. +- Antiquotations are evaluated in a context where there is exactly one goal +under focus, with the hypotheses coming from the current environment extended +with the bound variables of the term, and the resulting term is fed into the +quoted term. + +Relative orders of evaluation of antiquotations and quoted term is not +specified. + +## Patterns + +Terms can be used in pattern position just as any Ltac constructor. The accepted +syntax is a subset of the constr syntax in Ltac term position, where +antiquotations are variables binding in the right-hand side. + +Constructors and destructors can be derived from this. E.g. the previous +var-manipulating functions can be defined as follows. + +``` +let mkVar : ident -> constr = fun id -> << $ident:id$ >> + +let destVar : constr -> ident = function +| << $ident:x$ >> -> x +| _ -> fail () +``` + +One should be careful in patterns not to mix the syntax for evars with the one +for bound variables. + +The usual match construction from Ltac1 can be derived from those primitive +operations. We should provide syntactic sugar to do so. + +We need to decide how to handle bound variables in antiquotations, both in term +and pattern position. Should they bind? Should they not? What is the semantics +of the following snippet? + +``` +let foo = function << let x := t in $p$ >> -> p +let bar p = << let x := t in $p$ >> +``` + +What about the various kind of constrs? Untyped vs. typed, plus caring about +the context? + +### Lists and Gallina `match` + +It should be possible to manipulate Gallina `match` statements in a relatively +pain-free way. For this reason, there should be a way to match on lists: + +``` +let replace_args = function << $f$ $a1 .. an$ >> + << $g$ $b1 .. bn$ >> + -> << $f$ $b1 .. bn$ >> +let head = function << $f$ $a1 .. an$ >> -> << $f$ >> +let args : constr -> constr list = function << $f$ $a1 .. an$ >> -> [a1 ; .. ; an] +let apply (f : constr) : constr list -> constr = function +| $a1 .. an$ -> << $f$ $a1 .. an$ >> +let complicated_identity v = (let f = head v in let xs = args v in apply f xs) + +let open_term_under_binders = function << fun $a1 .. an$ => $body$ >> -> << $body$ >> +let binders : constr -> ident list = function << fun $a1 .. an$ => $body$ >> -> [a1 ; .. ; an] +let close_term (body : constr) : ident list -> constr = function $a1 .. an$ -> << fun $a1 .. an$ => $body$ >> +let complicated_function_identity v = + let b = open_term_under_binders v in + let xs = binders v in + close_term body xs +``` + +We could implement the `@?P` pattern as something like the desugaring rule: +``` +rule + (match term with + | (@?P a1 .. an)) + ~> + let P = type_check (<< fun $a1 .. an$ => $term$ >>) in ... +``` +The call to `type_check` ensures that there are no remaining holes in the term. +It is, perhaps, overkill. + +Then we could destructure a `match` via syntax like: +``` +let match_to_eta = function +| << match $t$ as $t'$ in $Ty$ return $P$ with + | $c1$ => $v1$ + .. + | $cm$ => $vm$ + end >> + -> << match $t$ in $Ty$ return $Ty$ with + | $c1$ => $c1$ + .. + | $cm$ => $cm$ + end >> +``` +which would take something like `match b with true => 0 | false => 1 end` and +return `match b with true => true | false => false end`. + +We should be able to construct the eliminators for inductive types +in Ltac 2.0, using this syntax to generate the bodies, together with some +primitives for acquiring the relevant types. + + +**Questions**: +- What exactly are the semantics of `..`? +- Should it be `$a1 .. an$` or `$a1$ .. $an$`? +- This syntax suggests that when open terms are used in binding positions, + unbound variables should become binding patterns. That is, if you have + `v` which has been constructed as `<< @cons _ $x$ $xs$ >>`, then + `<< fun ls : list nat => match ls with $v$ => $v$ | _ => nil end >>` should + be the eta-expansion of `ls`. Is this desired semantics? Are there issues + with it? + +# Notations + +Notations are the crux of the usability of Ltac. We should be able to recover +a feeling similar to the old implementation by using and abusing notations. +This would be done at at level totally different from the semantics, which +is not what is happening as of today. + +## Scopes + +We would like to attach some scopes to identifiers, so that it could be possible +to write e.g. + +``` +Ltac intro : string -> unit := ... + +Goal True -> True. +Proof. +intro "H". (** We require the quote here, as this is not a notation *) +Undo. +Top.intro "H". (** An alternative way, by fully qualifying the tactic *) +Abort. + +Tactic Notation "intro" ident(i) := intro i. + +Goal True -> True. +Proof. +intro H. +(** This sequence of tokens is elaborated at parsing time into [Top.intro "H"] + thanks to the above notation. *) +Undo. +Top.intro "H". +(** Here, the core tactic is still reachable using the fully qualified name *) +Abort. +``` + +A typical notation that would be useful is the Coq term one, so that the +following is possible. + +``` +Ltac destruct : constr -> unit := ... + +Tactic Notation "destruct" constr(x) := destruct x. + +Goal False -> True. +Proof. +intro H. (** assuming we have the previous notation in scope *) +destruct H. (** H is interpreted in the current goal? *) +Undo. +Top.destruct << H >> (** alternative without notation *) +``` + +Another one, probably useful for transition, would be a scope `legacy_constr` +that parses an identifier s.t. `legacy_constr(H)` elaborates to +`hyp H + mkVar H`. + +One should be able to define new scopes, by giving them a qualified name, +a old scope used for the parsing rule, and an expansion macro. We can maybe +unify such a scope creation process with the tactic notation one? + +## Syntactic sugar + +A few dedicated syntaxes should be built-in into Ltac2 for easy manipulation +of Coq-specific data. + +### Identifiers + +We need to write identifiers as easily as strings. What about `#foo` standing +for the identifier `foo`? + +### Hypotheses + +We need to be able to access easily a hypothesis from its name. What about +`` `foo `` being a shorthand for `hyp "foo"`? This needs to be accessible inside +terms as well. + +# Transition path + +TODO -- cgit v1.2.3 From 60e581f6fbcf033e134291016351492d9df7e319 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 1 Aug 2017 02:01:09 +0200 Subject: Fixup doc --- doc/ltac2.md | 2 ++ 1 file changed, 2 insertions(+) diff --git a/doc/ltac2.md b/doc/ltac2.md index 38a56d3d6b..e2aa4cfb3b 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -126,8 +126,10 @@ statically defined, but can instead by extended dynamically. A typical example is the standard `exn` type. Pattern-matching must always include a catch-all clause. They can be extended by the following command. +``` VERNAC ::= | "Ltac2" "Type" TYPEPARAMS QUALID ":=" "[" CONSTRUCTORDEF "]" +``` ## Term Syntax -- cgit v1.2.3 From 21087463e0a14bd101e01683c6dd7850fcccb395 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 1 Aug 2017 02:03:30 +0200 Subject: Fixup doc --- doc/ltac2.md | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/doc/ltac2.md b/doc/ltac2.md index e2aa4cfb3b..abd5493cec 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -263,6 +263,8 @@ is an evaluation context. ``` E[throw e] ≡ throw e + +(e value) ``` There is currently no way to catch such an exception and it is a design choice. @@ -286,7 +288,7 @@ If one sees thunks as lazy lists, then `zero` is the empty list and `plus` is list concatenation, while `case` is pattern-matching. The backtracking is first-class, i.e. one can write -`plus "x" (fun () -> "y") : string` producing a backtracking string. +`plus "x" (fun () => "y") : string` producing a backtracking string. These operations are expected to satisfy a few equations, most notably that they form a monoid compatible with sequentialization. @@ -299,7 +301,7 @@ plus (plus t f) g ≡ plus t (fun e -> plus (f e) g) case (fun () -> zero e) ≡ Err e case (fun () -> plus (fun () -> t) f) ≡ Val t f -let x := zero e in u ≡ fail e +let x := zero e in u ≡ zero e let x := plus t f in u ≡ plus (fun () -> let x := t in u) (fun e -> let x := f e in u) (t, u, f, g, e values) -- cgit v1.2.3 From 7cd31681eb5e3ccc7e7e920bb7eebe92827f6b16 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 1 Aug 2017 11:37:45 +0200 Subject: More in documentation. --- doc/ltac2.md | 385 ++++++++++++++++++++++++++--------------------------------- 1 file changed, 167 insertions(+), 218 deletions(-) diff --git a/doc/ltac2.md b/doc/ltac2.md index abd5493cec..fee19e5df0 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -185,6 +185,20 @@ There is a dedicated syntax for list and array litterals. Limitations: for now, deep pattern matching is not implemented yet. +## Ltac Definitions + +One can define a new global Ltac2 value using the following syntax. +``` +VERNAC ::= +| "Ltac2" RECFLAG LIDENT ":=" TERM +``` + +For semantic reasons, the body of the Ltac2 definition must be a syntactical +value, i.e. a function, a constant or a pure constructor recursively applied to +values. + +If the `RECFLAG` is set, the tactic is expanded into a recursive binding. + ## Reduction We use the usual ML call-by-value reduction, with an otherwise unspecified @@ -235,8 +249,8 @@ Intuitively a thunk of type `unit -> 'a` can do the following: - It can fail in a non-recoverable way - It can use first-class backtrack. The proper way to figure that is that we morally have the following isomorphism: - `(unit -> 'a) ~ unit -> ('a + (exn -> 'a))` i.e. thunks can produce a list - of results waiting for failure exceptions. + `(unit -> 'a) ~ (unit -> exn + ('a * (exn -> 'a)))` i.e. thunks can produce a + lazy list of results where each tail is waiting for a continuation exception. - It can access a backtracking proof state, made out amongst other things of the current evar assignation and the list of goals under focus. @@ -245,7 +259,12 @@ Intuitively a thunk of type `unit -> 'a` can do the following: The Ltac2 language features non-backtracking IO, notably mutable data and printing operations. -Mutable fields of records can be modified using the set syntax +Mutable fields of records can be modified using the set syntax. Likewise, +built-in types like `string` and `array` feature imperative assignment. See +modules `String` and `Array` respectively. + +A few printing primitives are provided in the `Message` module, allowing to +display information to the user. ### Fatal errors @@ -328,14 +347,7 @@ It is guaranteed that when evaluating `enter f`, `f` is called with exactly one goal under focus. Note that `f` may be called several times, or never, depending on the number of goals under focus before the call to `enter`. -A more expressive primitive allows to retrieve the data returned by each tactic -and store it in a list. - -``` -val enter_val : (unit -> 'a) -> 'a list -``` - -Accessing the goal data is then implicit in the Ltac2 primitives, and may fail +Accessing the goal data is then implicit in the Ltac2 primitives, and may panic if the invariants are not respected. The two essential functions for observing goals are given below. @@ -344,7 +356,7 @@ val hyp : ident -> constr val goal : unit -> constr ``` -The two above functions fail if there is not exactly one goal under focus. +The two above functions panic if there is not exactly one goal under focus. In addition, `hyp` may also fail if there is no hypothesis with the corresponding name. @@ -366,94 +378,106 @@ We should stop doing that! We need to mark when quoting and unquoting, although we need to do that in a short and elegant way so as not to be too cumbersome to the user. -## Syntax example - -Here is a suggestive example of a reasonable syntax. +## Generic Syntax for Quotations +In general, quotations can be introduced in term by the following syntax, where +`QUOTENTRY` is some parsing entry. ``` -let var := "H" in (* a string *) -let c := << fun $var$ => 0 >> (* the Coq term "fun H => 0" *) -let c' := << let x := $c$ in nat >> (* the Coq term "let x := fun H => 0 in nat" *) -... +TERM ::= +| QUOTNAME ":" "(" QUOTENTRY ")" + +QUOTNAME := IDENT ``` -## Term quotation +The current implementation recognizes the following built-in quotations: +- "ident", which parses identifiers (type `Init.ident`). +- "constr", which parses Coq terms and produces an-evar free term at runtime + (type `Init.constr`). +- "open_constr", which parses Coq terms and produces a term potentially with + holes at runtime (type `Init.constr` as well). +- "pattern", which parses Coq patterns and produces a pattern used for term + matching (type `Init.pattern`). -### Syntax +The following syntactic sugar is provided for two common cases. +- `@id` is the same as ident:(id) +- `'t` is the same as open_constr:(t) -It is better to define primitively the quoting syntax to build terms, as this -is more robust to changes. +## Term Antiquotations -``` -t, u ::= ... | << constr >> -``` +### Syntax -The `constr` datatype have the same syntax as the usual Coq -terms, except that it also allows antiquotations of the form `$t$` whose type -is statically inferred from the position, e.g. +One can also insert Ltac2 code into Coq term, similarly to what was possible in +Ltac1. ``` -<< let $t$ := $u$ >> (** [t] is an ident, [u] is a constr *) +COQCONSTR ::= +| "ltac2" ":" "(" TERM ")" ``` -As the term syntax implicitly allows to inject other classes without marking, -antiquotations can refer explicitly to which class they belong to overcome this -limitation. - -``` -<< $ident:t$ >> (** [t] is an ident, and the corresponding constr is [GVar t] *) -<< $ref:t$ >> (** [t] is a reference, and the corresponding constr is [GRef t] *) -``` +Antiquoted terms are expected to have type `unit`, as they are only evaluated +for their side-effects. ### Semantics -Interpretation of a quoted constr is done in two phases, internalization and +Interpretation of a quoted Coq term is done in two phases, internalization and evaluation. -- During internalization, variables are resolved and antiquotations are - type-checked as Ltac2 terms, effectively producing a `glob_constr` in Coq - implementation terminology, potentially ill-typed as a Coq term. -- During evaluation, a quoted term is fully evaluated to a kernel term, and is - in particular type-checked in the current environment. -Internalization is part of the static semantics, i.e. it is done at typing -time, while evaluation is part of the dynamic semantics, i.e. it is done when -a term gets effectively computed. +- Internalization is part of the static semantics, i.e. it is done at Ltac2 + typing time. +- Evaluation is part of the dynamic semantics, i.e. it is done when + a term gets effectively computed by Ltac2. #### Static semantics -The typing rule of a quoted constr is given below, where the `eᵢ` refer to -antiquoted terms. +During internalization, Coq variables are resolved and antiquotations are +type-checked as Ltac2 terms, effectively producing a `glob_constr` in Coq +implementation terminology. Note that although it went through the +type-checking of *Ltac2*, the resulting term has not been fully computed and +is potentially ill-typed as a Coq term. ``` - Γ ⊢ e₁ : unit Γ ⊢ eₙ : unit -==================================== - Γ ⊢ << c{$e₁$, ..., $eₙ$} >> : constr +Ltac2 Definition myconstr () := constr:(nat -> 0). +// Valid with type `unit -> constr`, but will fail at runtime. ``` -Note that the **static** environment of typing of antiquotations is **not** -expanded by the binders from the term. Namely, it means that the following +Term antiquotations are type-checked in the enclosing Ltac2 typing context +of the corresponding term expression. For instance, the following with +type-check. + +``` +let x := '0 in constr:(1 + ltac2:(exact x)) +// type constr +``` + +Beware that the typing environment of typing of antiquotations is **not** +expanded by the Coq binders from the term. Namely, it means that the following expression will **not** type-check. ``` -<< fun x : nat => $exact x$ >> +constr:(fun x : nat => ltac2:(exact x)) +// Error: Unbound variable 'x' ``` There is a simple reason for that, which is that the following expression would not make sense in general. ``` -<< fun x : nat => $clear x; exact x$ >> +constr:(fun x : nat => ltac2:(clear @x; exact x)) ``` -Rather, the tactic writer has to resort to the **dynamic** environment, and must -write instead something that amounts to the following. +Rather, the tactic writer has to resort to the **dynamic** goal environment, +and must write instead explicitly that she is accessing a hypothesis, typically +as follows. ``` -<< fun x : nat => $exact (hyp "x")$ >> +constr:(fun x : nat => ltac2:(hyp @x)) ``` -Obviously, we need to provide syntactic sugar to make this tractable. See the -corresponding section for more details. +The `ltac2:(hyp @x)` pattern is so common that we provide a dedicated Coq +term notation for it. #### Dynamic semantics +During evaluation, a quoted term is fully evaluated to a kernel term, and is +in particular type-checked in the current environment. + Evaluation of a quoted term is described below. - The quoted term is evaluated by the pretyper. - Antiquotations are evaluated in a context where there is exactly one goal @@ -464,183 +488,108 @@ quoted term. Relative orders of evaluation of antiquotations and quoted term is not specified. -## Patterns - -Terms can be used in pattern position just as any Ltac constructor. The accepted -syntax is a subset of the constr syntax in Ltac term position, where -antiquotations are variables binding in the right-hand side. - -Constructors and destructors can be derived from this. E.g. the previous -var-manipulating functions can be defined as follows. - +For instance, in the following example, `tac` will be evaluated in a context +with exactly one goal under focus, whose last hypothesis is `H : nat`. The +whole expression will thus evaluate to the term `fun H : nat => nat`. ``` -let mkVar : ident -> constr = fun id -> << $ident:id$ >> - -let destVar : constr -> ident = function -| << $ident:x$ >> -> x -| _ -> fail () +let tac () := hyp @H in constr:(fun H : nat => ltac2:(tac ())) ``` -One should be careful in patterns not to mix the syntax for evars with the one -for bound variables. - -The usual match construction from Ltac1 can be derived from those primitive -operations. We should provide syntactic sugar to do so. - -We need to decide how to handle bound variables in antiquotations, both in term -and pattern position. Should they bind? Should they not? What is the semantics -of the following snippet? - -``` -let foo = function << let x := t in $p$ >> -> p -let bar p = << let x := t in $p$ >> -``` +Many standard tactics perform type-checking of their argument before going +further. It is your duty to ensure that terms are well-typed when calling +such tactics. Failure to do so will result in non-recoverable exceptions. -What about the various kind of constrs? Untyped vs. typed, plus caring about -the context? - -### Lists and Gallina `match` - -It should be possible to manipulate Gallina `match` statements in a relatively -pain-free way. For this reason, there should be a way to match on lists: - -``` -let replace_args = function << $f$ $a1 .. an$ >> - << $g$ $b1 .. bn$ >> - -> << $f$ $b1 .. bn$ >> -let head = function << $f$ $a1 .. an$ >> -> << $f$ >> -let args : constr -> constr list = function << $f$ $a1 .. an$ >> -> [a1 ; .. ; an] -let apply (f : constr) : constr list -> constr = function -| $a1 .. an$ -> << $f$ $a1 .. an$ >> -let complicated_identity v = (let f = head v in let xs = args v in apply f xs) - -let open_term_under_binders = function << fun $a1 .. an$ => $body$ >> -> << $body$ >> -let binders : constr -> ident list = function << fun $a1 .. an$ => $body$ >> -> [a1 ; .. ; an] -let close_term (body : constr) : ident list -> constr = function $a1 .. an$ -> << fun $a1 .. an$ => $body$ >> -let complicated_function_identity v = - let b = open_term_under_binders v in - let xs = binders v in - close_term body xs -``` - -We could implement the `@?P` pattern as something like the desugaring rule: -``` -rule - (match term with - | (@?P a1 .. an)) - ~> - let P = type_check (<< fun $a1 .. an$ => $term$ >>) in ... -``` -The call to `type_check` ensures that there are no remaining holes in the term. -It is, perhaps, overkill. - -Then we could destructure a `match` via syntax like: -``` -let match_to_eta = function -| << match $t$ as $t'$ in $Ty$ return $P$ with - | $c1$ => $v1$ - .. - | $cm$ => $vm$ - end >> - -> << match $t$ in $Ty$ return $Ty$ with - | $c1$ => $c1$ - .. - | $cm$ => $cm$ - end >> -``` -which would take something like `match b with true => 0 | false => 1 end` and -return `match b with true => true | false => false end`. - -We should be able to construct the eliminators for inductive types -in Ltac 2.0, using this syntax to generate the bodies, together with some -primitives for acquiring the relevant types. +## Patterns +Terms can be used in pattern position just as any Ltac constructor. The accepted +syntax is a subset of the constr syntax in Ltac term position. It does not +allow antiquotations. -**Questions**: -- What exactly are the semantics of `..`? -- Should it be `$a1 .. an$` or `$a1$ .. $an$`? -- This syntax suggests that when open terms are used in binding positions, - unbound variables should become binding patterns. That is, if you have - `v` which has been constructed as `<< @cons _ $x$ $xs$ >>`, then - `<< fun ls : list nat => match ls with $v$ => $v$ | _ => nil end >>` should - be the eta-expansion of `ls`. Is this desired semantics? Are there issues - with it? +Patterns quotations are typically used with the matching functions provided +in the `Pattern` module. # Notations -Notations are the crux of the usability of Ltac. We should be able to recover +Notations are the crux of the usability of Ltac1. We should be able to recover a feeling similar to the old implementation by using and abusing notations. -This would be done at at level totally different from the semantics, which -is not what is happening as of today. ## Scopes -We would like to attach some scopes to identifiers, so that it could be possible -to write e.g. +A scope is a name given to a grammar entry used to produce some Ltac2 expression +at parsing time. Scopes are described using a form of S-expression. ``` -Ltac intro : string -> unit := ... - -Goal True -> True. -Proof. -intro "H". (** We require the quote here, as this is not a notation *) -Undo. -Top.intro "H". (** An alternative way, by fully qualifying the tactic *) -Abort. - -Tactic Notation "intro" ident(i) := intro i. - -Goal True -> True. -Proof. -intro H. -(** This sequence of tokens is elaborated at parsing time into [Top.intro "H"] - thanks to the above notation. *) -Undo. -Top.intro "H". -(** Here, the core tactic is still reachable using the fully qualified name *) -Abort. +SCOPE := +| STRING +| INT +| LIDENT ( "(" SCOPE₀ "," ... "," SCOPEₙ ")" ) +``` + +A few scopes contain antiquotation features. For sake of uniformity, all +antiquotations are introduced by the syntax `"$" VAR`. + +The following scopes are built-in. +- constr: + + parses `c = COQCONSTR` and produces `constr:(c)` +- ident: + + parses `id = IDENT` and produces `ident:(id)` + + parses `"$" (x = IDENT)` and produces the variable `x` +- list0(*scope*): + + if *scope* parses `ENTRY`, parses ̀`(x₀, ..., xₙ = ENTRY*)` and produces + `[x₀; ...; xₙ]`. +- list0(*scope*, sep = STRING): + + if *scope* parses `ENTRY`, parses `(x₀ = ENTRY, "sep", ..., "sep", xₙ = ENTRY)` + and produces `[x₀; ...; xₙ]`. +- list1: same as list0 (with or without separator) but parses `ENTRY+` instead + of `ENTRY*`. +- opt(*scope*) + + if *scope* parses `ENTRY`, parses `ENTRY?` and produces either `None` or + `Some x` where `x` is the parsed expression. +- self: + + parses a Ltac2 expression at the current level and return it as is. +- next: + + parses a Ltac2 expression at the next level and return it as is. +- tactic(n = INT): + + parses a Ltac2 expression at the provided level *n* and return it as is. +- thunk(*scope*): + parses the same as *scope*, and if *e* is the parsed expression, returns + `fun () => e`. + +For now there is no way to declare new scopes from Ltac2 side, but this is +planned. + +## Notations + +The Ltac2 parser can be extended by syntactic notations. ``` +VERNAC ::= +| "Ltac2" "Notation" TOKEN+ LEVEL? ":=" TERM -A typical notation that would be useful is the Coq term one, so that the -following is possible. +LEVEL := ":" INT +TOKEN := +| VAR "(" SCOPE ")" +| STRING ``` -Ltac destruct : constr -> unit := ... -Tactic Notation "destruct" constr(x) := destruct x. +A Ltac2 notation adds a parsing rule to the Ltac2 grammar, which is expanded +to the provided body where every token from the notation is let-bound to the +corresponding generated expression. -Goal False -> True. -Proof. -intro H. (** assuming we have the previous notation in scope *) -destruct H. (** H is interpreted in the current goal? *) -Undo. -Top.destruct << H >> (** alternative without notation *) +For instance, assume we perform: +``` +Ltac2 Notation "foo" c(thunk(constr)) ids(list0(ident)) := Bar.f c ids. +``` +Then the following expression +``` +let y := @X in foo (nat -> nat) x ?y +``` +will expand at parsing time to +``` +let y := @X in +let c := fun () => constr:(nat -> nat) with ids := [@x; y] in Bar.f c ids ``` -Another one, probably useful for transition, would be a scope `legacy_constr` -that parses an identifier s.t. `legacy_constr(H)` elaborates to -`hyp H + mkVar H`. - -One should be able to define new scopes, by giving them a qualified name, -a old scope used for the parsing rule, and an expansion macro. We can maybe -unify such a scope creation process with the tactic notation one? - -## Syntactic sugar - -A few dedicated syntaxes should be built-in into Ltac2 for easy manipulation -of Coq-specific data. - -### Identifiers - -We need to write identifiers as easily as strings. What about `#foo` standing -for the identifier `foo`? - -### Hypotheses - -We need to be able to access easily a hypothesis from its name. What about -`` `foo `` being a shorthand for `hyp "foo"`? This needs to be accessible inside -terms as well. - -# Transition path - -TODO +Beware that the order of evaluation of multiple let-bindings is not specified, +so that you may have to resort to thunking to ensure that side-effects are +performed at the right time. -- cgit v1.2.3 From 8a8fd265158fa3fe7eea65b50c3da722e81fa688 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 1 Aug 2017 14:43:30 +0200 Subject: Binding more primitive tactics. --- src/tac2core.ml | 2 +- src/tac2ffi.ml | 3 ++ src/tac2ffi.mli | 3 ++ src/tac2stdlib.ml | 80 +++++++++++++++++++++++++++++++++++++++++++++++++++++ tests/stuff/ltac2.v | 8 ++++++ theories/Std.v | 25 +++++++++++++++++ 6 files changed, 120 insertions(+), 1 deletion(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index e0a65dde2d..fef16dcc06 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -325,7 +325,7 @@ let prm_constr_kind : ml_tactic = function |]) | Const (cst, u) -> ValBlk (10, [| - Value.of_ext Value.val_constant cst; + Value.of_constant cst; of_instance sigma u; |]) | Ind (ind, u) -> diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index 74e2b02aeb..49b49d92fd 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -122,3 +122,6 @@ let of_array f vl = ValBlk (0, Array.map f vl) let to_array f = function | ValBlk (0, vl) -> Array.map f vl | _ -> assert false + +let of_constant c = of_ext val_constant c +let to_constant c = to_ext val_constant c diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index 3f429995ce..b69ca9a382 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -59,6 +59,9 @@ val to_pattern : valexpr -> Pattern.constr_pattern val of_pp : Pp.t -> valexpr val to_pp : valexpr -> Pp.t +val of_constant : Constant.t -> valexpr +val to_constant : valexpr -> Constant.t + val of_ext : 'a Val.typ -> 'a -> valexpr val to_ext : 'a Val.typ -> valexpr -> 'a diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 7c7b539113..e093b5c97f 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -9,6 +9,7 @@ open Names open Locus open Misctypes +open Genredexpr open Tac2expr open Tac2core open Proofview.Notations @@ -66,6 +67,24 @@ let to_clause = function { onhyps = hyps; concl_occs = to_occurrences to_int_or_var concl; } | _ -> assert false +let to_evaluable_ref = function +| ValBlk (0, [| id |]) -> EvalVarRef (Value.to_ident id) +| ValBlk (1, [| cst |]) -> EvalConstRef (Value.to_constant cst) +| _ -> assert false + +let to_red_flag = function +| ValBlk (0, [| beta; iota; fix; cofix; zeta; delta; const |]) -> + { + rBeta = Value.to_bool beta; + rMatch = Value.to_bool iota; + rFix = Value.to_bool fix; + rCofix = Value.to_bool cofix; + rZeta = Value.to_bool zeta; + rDelta = Value.to_bool delta; + rConst = Value.to_list to_evaluable_ref const; + } +| _ -> assert false + (** Standard tactics sharing their implementation with Ltac1 *) let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } @@ -81,6 +100,8 @@ let wrap f = let wrap_unit f = return () >>= fun () -> f (); return v_unit +let thaw f = Tac2interp.interp_app f [v_unit] + let define_prim0 name tac = let tac = function | [_] -> lift tac @@ -102,6 +123,13 @@ let define_prim2 name tac = in Tac2env.define_primitive (pname name) tac +let define_prim3 name tac = + let tac = function + | [x; y; z] -> lift (tac x y z) + | _ -> assert false + in + Tac2env.define_primitive (pname name) tac + (** Tactics from Tacexpr *) let () = define_prim2 "tac_eelim" begin fun c copt -> @@ -125,6 +153,58 @@ let () = define_prim1 "tac_egeneralize" begin fun cl -> Tactics.new_generalize_gen cl end +let () = define_prim2 "tac_pose" begin fun idopt c -> + let na = to_name idopt in + let c = Value.to_constr c in + Tactics.letin_tac None na c None Locusops.nowhere +end + +let () = define_prim3 "tac_set" begin fun idopt c cl -> + let na = to_name idopt in + let cl = to_clause cl in + Proofview.tclEVARMAP >>= fun sigma -> + thaw c >>= fun c -> + let c = Value.to_constr c in + Tactics.letin_pat_tac false None na (sigma, c) cl +end + +let () = define_prim3 "tac_eset" begin fun idopt c cl -> + let na = to_name idopt in + let cl = to_clause cl in + Proofview.tclEVARMAP >>= fun sigma -> + thaw c >>= fun c -> + let c = Value.to_constr c in + Tactics.letin_pat_tac true None na (sigma, c) cl +end + +let () = define_prim1 "tac_red" begin fun cl -> + let cl = to_clause cl in + Tactics.reduce (Red false) cl +end + +let () = define_prim1 "tac_hnf" begin fun cl -> + let cl = to_clause cl in + Tactics.reduce Hnf cl +end + +let () = define_prim2 "tac_cbv" begin fun flags cl -> + let flags = to_red_flag flags in + let cl = to_clause cl in + Tactics.reduce (Cbv flags) cl +end + +let () = define_prim2 "tac_cbn" begin fun flags cl -> + let flags = to_red_flag flags in + let cl = to_clause cl in + Tactics.reduce (Cbn flags) cl +end + +let () = define_prim2 "tac_lazy" begin fun flags cl -> + let flags = to_red_flag flags in + let cl = to_clause cl in + Tactics.reduce (Lazy flags) cl +end + (** Tactics from coretactics *) let () = define_prim0 "tac_reflexivity" Tactics.intros_reflexivity diff --git a/tests/stuff/ltac2.v b/tests/stuff/ltac2.v index ece6fca06a..6b30d42c09 100644 --- a/tests/stuff/ltac2.v +++ b/tests/stuff/ltac2.v @@ -143,3 +143,11 @@ Ltac2 rec do n tac := match Int.equal n 0 with end. Print Ltac2 do. + +Goal forall x, 1 + x = x + 1. +Proof. +refine (fun () => '(fun x => _)). +Std.cbv { + Std.rBeta := true; Std.rMatch := true; Std.rFix := true; Std.rCofix := true; + Std.rZeta := true; Std.rDelta := false; rConst := []; +} { Std.on_hyps := None; Std.on_concl := Std.AllOccurrences }. diff --git a/theories/Std.v b/theories/Std.v index a9eced6cbb..3070c2e005 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -34,6 +34,21 @@ Ltac2 Type clause := { on_concl : occurrences; }. +Ltac2 Type evaluable_reference := [ +| EvalVarRef (ident) +| EvalConstRef (constant) +]. + +Ltac2 Type red_flags := { + rBeta : bool; + rMatch : bool; + rFix : bool; + rCofix : bool; + rZeta : bool; + rDelta : bool; (** true = delta all but rConst; false = delta only on rConst*) + rConst : evaluable_reference list +}. + (** Standard, built-in tactics. See Ltac1 for documentation. *) Ltac2 @ external eelim : constr_with_bindings -> constr_with_bindings option -> unit := "ltac2" "tac_eelim". @@ -41,6 +56,16 @@ Ltac2 @ external ecase : constr_with_bindings -> unit := "ltac2" "tac_ecase". Ltac2 @ external egeneralize : (constr * occurrences * ident option) list -> unit := "ltac2" "tac_egeneralize". +Ltac2 @ external pose : ident option -> constr -> unit := "ltac2" "tac_pose". +Ltac2 @ external set : ident option -> (unit -> constr) -> clause -> unit := "ltac2" "tac_set". +Ltac2 @ external eset : ident option -> (unit -> constr) -> clause -> unit := "ltac2" "tac_eset". + +Ltac2 @ external red : clause -> unit := "ltac2" "tac_red". +Ltac2 @ external hnf : clause -> unit := "ltac2" "tac_hnf". +Ltac2 @ external cbv : red_flags -> clause -> unit := "ltac2" "tac_cbv". +Ltac2 @ external cbn : red_flags -> clause -> unit := "ltac2" "tac_cbn". +Ltac2 @ external lazy : red_flags -> clause -> unit := "ltac2" "tac_lazy". + Ltac2 @ external reflexivity : unit -> unit := "ltac2" "tac_reflexivity". Ltac2 @ external assumption : unit -> unit := "ltac2" "tac_assumption". -- cgit v1.2.3 From 30fc910b01f61ce3691ed63a0908c1c60cee76dd Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 1 Aug 2017 16:46:29 +0200 Subject: Fix documentation. --- doc/ltac2.md | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/doc/ltac2.md b/doc/ltac2.md index fee19e5df0..7a3b2181f8 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -467,11 +467,11 @@ Rather, the tactic writer has to resort to the **dynamic** goal environment, and must write instead explicitly that she is accessing a hypothesis, typically as follows. ``` -constr:(fun x : nat => ltac2:(hyp @x)) +constr:(fun x : nat => ltac2:(exact (hyp @x))) ``` -The `ltac2:(hyp @x)` pattern is so common that we provide a dedicated Coq -term notation for it. +The `ltac2:(exact (hyp @x))` pattern is so common that we provide dedicated +Ltac2 and Coq term notations for it. #### Dynamic semantics -- cgit v1.2.3 From c3be78f96b91a042944f9bee66bf0ea8d929a37d Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 1 Aug 2017 16:56:27 +0200 Subject: Introducing the all-mighty intro-patterns. --- _CoqProject | 1 + src/g_ltac2.ml4 | 73 ++++++++++++++++++++++++++++++++++++++++++++++------- src/tac2core.ml | 22 ++++++++++------ src/tac2core.mli | 6 ++--- src/tac2entries.ml | 1 + src/tac2entries.mli | 1 + src/tac2qexpr.mli | 36 ++++++++++++++++++++++++++ src/tac2quote.ml | 57 ++++++++++++++++++++++++++++++++++++++--- src/tac2quote.mli | 7 +++++ src/tac2stdlib.ml | 43 +++++++++++++++++++++++++++++++ theories/Std.v | 27 ++++++++++++++++++++ 11 files changed, 250 insertions(+), 24 deletions(-) create mode 100644 src/tac2qexpr.mli diff --git a/_CoqProject b/_CoqProject index ab73af1295..b8064c46a4 100644 --- a/_CoqProject +++ b/_CoqProject @@ -17,6 +17,7 @@ src/tac2ffi.ml src/tac2ffi.mli src/tac2core.ml src/tac2core.mli +src/tac2qexpr.mli src/tac2quote.ml src/tac2quote.mli src/tac2stdlib.ml diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 4a2f615df9..b058680645 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -14,6 +14,7 @@ open Pcoq open Constrexpr open Misctypes open Tac2expr +open Tac2qexpr open Ltac_plugin let err () = raise Stream.Failure @@ -44,9 +45,6 @@ let inj_wit wit loc x = CTacExt (loc, Genarg.in_gen (Genarg.rawwit wit) x) let inj_open_constr loc c = inj_wit Stdarg.wit_open_constr loc c let inj_pattern loc c = inj_wit Tac2env.wit_pattern loc c -let mk_constr ~loc kn args = - CTacApp (loc, CTacCst (loc, AbsKn (Other kn)), args) - let pattern_of_qualid loc id = if Tac2env.is_constructor (snd id) then CPatRef (loc, RelId id, []) else @@ -278,14 +276,19 @@ END open Tac2entries.Pltac +let loc_of_ne_list l = Loc.merge_opt (fst (List.hd l)) (fst (List.last l)) + GEXTEND Gram - GLOBAL: q_ident q_bindings; - q_ident: - [ [ id = Prim.ident -> Tac2quote.of_ident ~loc:!@loc id - | "$"; id = Prim.ident -> Tac2quote.of_variable ~loc:!@loc id + GLOBAL: q_ident q_bindings q_intropatterns; + ident_or_anti: + [ [ id = Prim.ident -> QExpr id + | "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] ; - simple_binding: + q_ident: + [ [ id = ident_or_anti -> Tac2quote.of_anti ~loc:!@loc Tac2quote.of_ident id ] ] + ; + simple_binding: [ [ "("; id = Prim.ident; ":="; c = Constr.lconstr; ")" -> Loc.tag ~loc:!@loc (NamedHyp id, Tac2quote.of_constr ~loc:!@loc c) | "("; n = Prim.natural; ":="; c = Constr.lconstr; ")" -> @@ -302,9 +305,61 @@ GEXTEND Gram ; q_bindings: [ [ "with"; bl = bindings -> bl - | -> mk_constr ~loc:!@loc Tac2core.Core.c_no_bindings [] + | -> Tac2quote.of_bindings ~loc:!@loc Misctypes.NoBindings ] ] ; + intropatterns: + [ [ l = LIST0 nonsimple_intropattern -> l ]] + ; +(* ne_intropatterns: *) +(* [ [ l = LIST1 nonsimple_intropattern -> l ]] *) +(* ; *) + or_and_intropattern: + [ [ "["; tc = LIST1 intropatterns SEP "|"; "]" -> QIntroOrPattern tc + | "()" -> QIntroAndPattern [] + | "("; si = simple_intropattern; ")" -> QIntroAndPattern [si] + | "("; si = simple_intropattern; ","; + tc = LIST1 simple_intropattern SEP "," ; ")" -> + QIntroAndPattern (si::tc) + | "("; si = simple_intropattern; "&"; + tc = LIST1 simple_intropattern SEP "&" ; ")" -> + (* (A & B & C) is translated into (A,(B,C)) *) + let rec pairify = function + | ([]|[_]|[_;_]) as l -> l + | t::q -> [t; (QIntroAction (QIntroOrAndPattern (QIntroAndPattern (pairify q))))] + in QIntroAndPattern (pairify (si::tc)) ] ] + ; + equality_intropattern: + [ [ "->" -> QIntroRewrite true + | "<-" -> QIntroRewrite false + | "[="; tc = intropatterns; "]" -> QIntroInjection tc ] ] + ; + naming_intropattern: + [ [ LEFTQMARK; prefix = ident_or_anti -> QIntroFresh prefix + | LEFTQMARK -> QIntroAnonymous + | id = ident_or_anti -> QIntroIdentifier id ] ] + ; + nonsimple_intropattern: + [ [ l = simple_intropattern -> l + | "*" -> QIntroForthcoming true + | "**" -> QIntroForthcoming false ]] + ; + simple_intropattern: + [ [ pat = simple_intropattern_closed -> +(* l = LIST0 ["%"; c = operconstr LEVEL "0" -> c] -> *) + (** TODO: handle %pat *) + pat + ] ] + ; + simple_intropattern_closed: + [ [ pat = or_and_intropattern -> QIntroAction (QIntroOrAndPattern pat) + | pat = equality_intropattern -> QIntroAction pat + | "_" -> QIntroAction QIntroWildcard + | pat = naming_intropattern -> QIntroNaming pat ] ] + ; + q_intropatterns: + [ [ ipat = intropatterns -> Tac2quote.of_intro_patterns ~loc:!@loc ipat ] ] + ; END (** Extension of constr syntax *) diff --git a/src/tac2core.ml b/src/tac2core.ml index fef16dcc06..266e3b5f11 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -42,10 +42,8 @@ let c_cons = coq_core "::" let c_none = coq_core "None" let c_some = coq_core "Some" -let t_bindings = std_core "bindings" -let c_no_bindings = std_core "NoBindings" -let c_implicit_bindings = std_core "ImplicitBindings" -let c_explicit_bindings = std_core "ExplicitBindings" +let c_true = coq_core "true" +let c_false = coq_core "false" let t_qhyp = std_core "hypothesis" let c_named_hyp = std_core "NamedHyp" @@ -853,6 +851,14 @@ let () = add_scope "ident" begin function | _ -> scope_fail () end +let () = add_scope "thunk" begin function +| [tok] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let act e = rthunk (act e) in + Tac2entries.ScopeRule (scope, act) +| _ -> scope_fail () +end + let () = add_scope "bindings" begin function | [] -> let scope = Extend.Aentry Tac2entries.Pltac.q_bindings in @@ -861,10 +867,10 @@ let () = add_scope "bindings" begin function | _ -> scope_fail () end -let () = add_scope "thunk" begin function -| [tok] -> - let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in - let act e = rthunk (act e) in +let () = add_scope "intropatterns" begin function +| [] -> + let scope = Extend.Aentry Tac2entries.Pltac.q_intropatterns in + let act tac = tac in Tac2entries.ScopeRule (scope, act) | _ -> scope_fail () end diff --git a/src/tac2core.mli b/src/tac2core.mli index 118b7aaa42..6fd48e85f7 100644 --- a/src/tac2core.mli +++ b/src/tac2core.mli @@ -24,10 +24,8 @@ val t_option : type_constant val t_string : type_constant val t_array : type_constant -val t_bindings : type_constant -val c_no_bindings : ltac_constructor -val c_implicit_bindings : ltac_constant -val c_explicit_bindings : ltac_constant +val c_true : ltac_constructor +val c_false : ltac_constructor val t_qhyp : type_constant val c_anon_hyp : ltac_constructor diff --git a/src/tac2entries.ml b/src/tac2entries.ml index d293a87975..d7ee07e9e2 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -26,6 +26,7 @@ let tac2expr = Pcoq.Gram.entry_create "tactic:tac2expr" let q_ident = Pcoq.Gram.entry_create "tactic:q_ident" let q_bindings = Pcoq.Gram.entry_create "tactic:q_bindings" +let q_intropatterns = Pcoq.Gram.entry_create "tactic:q_intropatterns" end (** Tactic definition *) diff --git a/src/tac2entries.mli b/src/tac2entries.mli index 4d5a234daf..e5031fdba2 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -59,4 +59,5 @@ val tac2expr : raw_tacexpr Pcoq.Gram.entry val q_ident : raw_tacexpr Pcoq.Gram.entry val q_bindings : raw_tacexpr Pcoq.Gram.entry +val q_intropatterns : raw_tacexpr Pcoq.Gram.entry end diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli new file mode 100644 index 0000000000..794281cc75 --- /dev/null +++ b/src/tac2qexpr.mli @@ -0,0 +1,36 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* f ?loc x +| QAnti (loc, id) -> of_variable ?loc id + let of_ident ?loc id = inj_wit ?loc Stdarg.wit_ident id let of_constr ?loc c = inj_wit ?loc Stdarg.wit_constr c +let of_bool ?loc b = + let c = if b then Core.c_true else Core.c_false in + constructor ?loc c [] + let rec of_list ?loc = function | [] -> constructor Core.c_nil [] | e :: l -> @@ -55,9 +70,45 @@ let of_qhyp ?loc = function let of_bindings ?loc = function | NoBindings -> - constructor ?loc Core.c_no_bindings [] + std_constructor ?loc "NoBindings" [] | ImplicitBindings tl -> - constructor ?loc Core.c_implicit_bindings [of_list ?loc tl] + std_constructor ?loc "ImplicitBindings" [of_list ?loc tl] | ExplicitBindings tl -> let tl = List.map (fun (loc, (qhyp, e)) -> of_pair ?loc (of_qhyp ?loc qhyp, e)) tl in - constructor ?loc Core.c_explicit_bindings [of_list ?loc tl] + std_constructor ?loc "ExplicitBindings" [of_list ?loc tl] + +let rec of_intro_pattern ?loc = function +| QIntroForthcoming b -> + std_constructor ?loc "IntroForthcoming" [of_bool b] +| QIntroNaming iname -> + std_constructor ?loc "IntroNaming" [of_intro_pattern_naming iname] +| QIntroAction iact -> + std_constructor ?loc "IntroAction" [of_intro_pattern_action iact] + +and of_intro_pattern_naming ?loc = function +| QIntroIdentifier id -> + std_constructor ?loc "IntroIdentifier" [of_anti ?loc of_ident id] +| QIntroFresh id -> + std_constructor ?loc "IntroFresh" [of_anti ?loc of_ident id] +| QIntroAnonymous -> + std_constructor ?loc "IntroAnonymous" [] + +and of_intro_pattern_action ?loc = function +| QIntroWildcard -> + std_constructor ?loc "IntroWildcard" [] +| QIntroOrAndPattern pat -> + std_constructor ?loc "IntroOrAndPattern" [of_or_and_intro_pattern ?loc pat] +| QIntroInjection il -> + std_constructor ?loc "IntroInjection" [of_intro_patterns ?loc il] +| QIntroRewrite b -> + std_constructor ?loc "IntroRewrite" [of_bool ?loc b] + +and of_or_and_intro_pattern ?loc = function +| QIntroOrPattern ill -> + let ill = List.map (of_intro_patterns ?loc) ill in + std_constructor ?loc "IntroOrPattern" [of_list ?loc ill] +| QIntroAndPattern il -> + std_constructor ?loc "IntroAndPattern" [of_intro_patterns ?loc il] + +and of_intro_patterns ?loc l = + of_list ?loc (List.map (of_intro_pattern ?loc) l) diff --git a/src/tac2quote.mli b/src/tac2quote.mli index ba6a878d50..32973ff5ba 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -8,6 +8,7 @@ open Names open Misctypes +open Tac2qexpr open Tac2expr (** Syntactic quoting of expressions. *) @@ -17,6 +18,8 @@ open Tac2expr val constructor : ?loc:Loc.t -> ltac_constructor -> raw_tacexpr list -> raw_tacexpr +val of_anti : ?loc:Loc.t -> (?loc:Loc.t -> 'a -> raw_tacexpr) -> 'a or_anti -> raw_tacexpr + val of_int : ?loc:Loc.t -> int -> raw_tacexpr val of_pair : ?loc:Loc.t -> raw_tacexpr * raw_tacexpr -> raw_tacexpr @@ -30,3 +33,7 @@ val of_constr : ?loc:Loc.t -> Constrexpr.constr_expr -> raw_tacexpr val of_list : ?loc:Loc.t -> raw_tacexpr list -> raw_tacexpr val of_bindings : ?loc:Loc.t -> raw_tacexpr bindings -> raw_tacexpr + +val of_intro_pattern : ?loc:Loc.t -> intro_pattern -> raw_tacexpr + +val of_intro_patterns : ?loc:Loc.t -> intro_pattern list -> raw_tacexpr diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index e093b5c97f..44fad48955 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -85,6 +85,39 @@ let to_red_flag = function } | _ -> assert false +let rec to_intro_pattern = function +| ValBlk (0, [| b |]) -> IntroForthcoming (Value.to_bool b) +| ValBlk (1, [| pat |]) -> IntroNaming (to_intro_pattern_naming pat) +| ValBlk (2, [| act |]) -> IntroAction (to_intro_pattern_action act) +| _ -> assert false + +and to_intro_pattern_naming = function +| ValBlk (0, [| id |]) -> IntroIdentifier (Value.to_ident id) +| ValBlk (1, [| id |]) -> IntroFresh (Value.to_ident id) +| ValInt 0 -> IntroAnonymous +| _ -> assert false + +and to_intro_pattern_action = function +| ValInt 0 -> IntroWildcard +| ValBlk (0, [| op |]) -> IntroOrAndPattern (to_or_and_intro_pattern op) +| ValBlk (1, [| inj |]) -> + let map ipat = Loc.tag (to_intro_pattern ipat) in + IntroInjection (Value.to_list map inj) +| ValBlk (2, [| _ |]) -> IntroApplyOn (assert false, assert false) (** TODO *) +| ValBlk (3, [| b |]) -> IntroRewrite (Value.to_bool b) +| _ -> assert false + +and to_or_and_intro_pattern = function +| ValBlk (0, [| ill |]) -> + IntroOrPattern (Value.to_list to_intro_patterns ill) +| ValBlk (1, [| il |]) -> + IntroAndPattern (to_intro_patterns il) +| _ -> assert false + +and to_intro_patterns il = + let map ipat = Loc.tag (to_intro_pattern ipat) in + Value.to_list map il + (** Standard tactics sharing their implementation with Ltac1 *) let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } @@ -132,6 +165,16 @@ let define_prim3 name tac = (** Tactics from Tacexpr *) +let () = define_prim1 "tac_intros" begin fun ipat -> + let ipat = to_intro_patterns ipat in + Tactics.intros_patterns false ipat +end + +let () = define_prim1 "tac_eintros" begin fun ipat -> + let ipat = to_intro_patterns ipat in + Tactics.intros_patterns true ipat +end + let () = define_prim2 "tac_eelim" begin fun c copt -> let c = to_constr_with_bindings c in let copt = Value.to_option to_constr_with_bindings copt in diff --git a/theories/Std.v b/theories/Std.v index 3070c2e005..a27790c35d 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -49,8 +49,35 @@ Ltac2 Type red_flags := { rConst : evaluable_reference list }. +Ltac2 Type 'a not_implemented. + +Ltac2 Type rec intro_pattern := [ +| IntroForthcoming (bool) +| IntroNaming (intro_pattern_naming) +| IntroAction (intro_pattern_action) +] +with intro_pattern_naming := [ +| IntroIdentifier (ident) +| IntroFresh (ident) +| IntroAnonymous +] +with intro_pattern_action := [ +| IntroWildcard +| IntroOrAndPattern (or_and_intro_pattern) +| IntroInjection (intro_pattern list) +| IntroApplyOn ((constr * intro_pattern) not_implemented) (* Not Implemented yet *) +| IntroRewrite (bool) +] +with or_and_intro_pattern := [ +| IntroOrPattern (intro_pattern list list) +| IntroAndPattern (intro_pattern list) +]. + (** Standard, built-in tactics. See Ltac1 for documentation. *) +Ltac2 @ external intros : intro_pattern list -> unit := "ltac2" "tac_intros". +Ltac2 @ external eintros : intro_pattern list -> unit := "ltac2" "tac_eintros". + Ltac2 @ external eelim : constr_with_bindings -> constr_with_bindings option -> unit := "ltac2" "tac_eelim". Ltac2 @ external ecase : constr_with_bindings -> unit := "ltac2" "tac_ecase". -- cgit v1.2.3 From 73ecd7e2f0136234f73f405a569858f2b0ecee9b Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 1 Aug 2017 16:56:27 +0200 Subject: Fix parsing of fresh ident antiquotations. --- src/g_ltac2.ml4 | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index b058680645..a09e99aa6b 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -335,8 +335,9 @@ GEXTEND Gram | "[="; tc = intropatterns; "]" -> QIntroInjection tc ] ] ; naming_intropattern: - [ [ LEFTQMARK; prefix = ident_or_anti -> QIntroFresh prefix - | LEFTQMARK -> QIntroAnonymous + [ [ LEFTQMARK; id = Prim.ident -> QIntroFresh (QExpr id) + | "?$"; id = Prim.ident -> QIntroFresh (QAnti (Loc.tag ~loc:!@loc id)) + | "?" -> QIntroAnonymous | id = ident_or_anti -> QIntroIdentifier id ] ] ; nonsimple_intropattern: -- cgit v1.2.3 From 6d8b31504efce96ec6d3011763ced0c631cf576a Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 1 Aug 2017 20:35:19 +0200 Subject: Don't reuse Coq AST for binding quotations. This allows antiquotations in binding lists. --- src/g_ltac2.ml4 | 27 ++++++++++++++++++--------- src/tac2qexpr.mli | 6 ++++++ src/tac2quote.ml | 8 ++++---- src/tac2quote.mli | 2 +- tests/example2.v | 6 ++++++ 5 files changed, 35 insertions(+), 14 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index a09e99aa6b..71fb59acf8 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -19,7 +19,7 @@ open Ltac_plugin let err () = raise Stream.Failure -(* idem for (x:=t) and (1:=t) *) +(* lookahead for (x:=t), (?x:=t) and (1:=t) *) let test_lpar_idnum_coloneq = Gram.Entry.of_parser "test_lpar_idnum_coloneq" (fun strm -> @@ -30,6 +30,13 @@ let test_lpar_idnum_coloneq = (match stream_nth 2 strm with | KEYWORD ":=" -> () | _ -> err ()) + | LEFTQMARK -> + (match stream_nth 2 strm with + | IDENT _ -> + (match stream_nth 3 strm with + | KEYWORD ":=" -> () + | _ -> err ()) + | _ -> err ()) | _ -> err ()) | _ -> err ()) @@ -282,30 +289,32 @@ GEXTEND Gram GLOBAL: q_ident q_bindings q_intropatterns; ident_or_anti: [ [ id = Prim.ident -> QExpr id - | "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) + | LEFTQMARK; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] ; q_ident: [ [ id = ident_or_anti -> Tac2quote.of_anti ~loc:!@loc Tac2quote.of_ident id ] ] ; simple_binding: - [ [ "("; id = Prim.ident; ":="; c = Constr.lconstr; ")" -> - Loc.tag ~loc:!@loc (NamedHyp id, Tac2quote.of_constr ~loc:!@loc c) + [ [ "("; LEFTQMARK; id = Prim.ident; ":="; c = Constr.lconstr; ")" -> + Loc.tag ~loc:!@loc (QAnti (Loc.tag ~loc:!@loc id), Tac2quote.of_constr ~loc:!@loc c) | "("; n = Prim.natural; ":="; c = Constr.lconstr; ")" -> - Loc.tag ~loc:!@loc (AnonHyp n, Tac2quote.of_constr ~loc:!@loc c) + Loc.tag ~loc:!@loc (QExpr (AnonHyp n), Tac2quote.of_constr ~loc:!@loc c) + | "("; id = Prim.ident; ":="; c = Constr.lconstr; ")" -> + Loc.tag ~loc:!@loc (QExpr (NamedHyp id), Tac2quote.of_constr ~loc:!@loc c) ] ] ; bindings: [ [ test_lpar_idnum_coloneq; bl = LIST1 simple_binding -> - Tac2quote.of_bindings ~loc:!@loc (ExplicitBindings bl) + QExplicitBindings bl | bl = LIST1 Constr.constr -> let bl = List.map (fun c -> Tac2quote.of_constr ~loc:!@loc c) bl in - Tac2quote.of_bindings ~loc:!@loc (Misctypes.ImplicitBindings bl) + QImplicitBindings bl ] ] ; q_bindings: - [ [ "with"; bl = bindings -> bl - | -> Tac2quote.of_bindings ~loc:!@loc Misctypes.NoBindings + [ [ "with"; bl = bindings -> Tac2quote.of_bindings ~loc:!@loc bl + | -> Tac2quote.of_bindings ~loc:!@loc QNoBindings ] ] ; intropatterns: diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index 794281cc75..b68efe73ac 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -9,6 +9,7 @@ open Util open Loc open Names +open Tac2expr (** Quoted variants of Ltac syntactic categories. Contrarily to the former, they sometimes allow anti-quotations. Used for notation scopes. *) @@ -17,6 +18,11 @@ type 'a or_anti = | QExpr of 'a | QAnti of Id.t located +type bindings = +| QImplicitBindings of raw_tacexpr list +| QExplicitBindings of (Misctypes.quantified_hypothesis or_anti * raw_tacexpr) Loc.located list +| QNoBindings + type intro_pattern = | QIntroForthcoming of bool | QIntroNaming of intro_pattern_naming diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 96a3a5d9b8..488bcb5201 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -69,12 +69,12 @@ let of_qhyp ?loc = function | NamedHyp id -> constructor Core.c_named_hyp [of_ident ?loc id] let of_bindings ?loc = function -| NoBindings -> +| QNoBindings -> std_constructor ?loc "NoBindings" [] -| ImplicitBindings tl -> +| QImplicitBindings tl -> std_constructor ?loc "ImplicitBindings" [of_list ?loc tl] -| ExplicitBindings tl -> - let tl = List.map (fun (loc, (qhyp, e)) -> of_pair ?loc (of_qhyp ?loc qhyp, e)) tl in +| QExplicitBindings tl -> + let tl = List.map (fun (loc, (qhyp, e)) -> of_pair ?loc (of_anti ?loc of_qhyp qhyp, e)) tl in std_constructor ?loc "ExplicitBindings" [of_list ?loc tl] let rec of_intro_pattern ?loc = function diff --git a/src/tac2quote.mli b/src/tac2quote.mli index 32973ff5ba..c9ee270d38 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -32,7 +32,7 @@ val of_constr : ?loc:Loc.t -> Constrexpr.constr_expr -> raw_tacexpr val of_list : ?loc:Loc.t -> raw_tacexpr list -> raw_tacexpr -val of_bindings : ?loc:Loc.t -> raw_tacexpr bindings -> raw_tacexpr +val of_bindings : ?loc:Loc.t -> bindings -> raw_tacexpr val of_intro_pattern : ?loc:Loc.t -> intro_pattern -> raw_tacexpr diff --git a/tests/example2.v b/tests/example2.v index 5efbf90b34..ffdb723ffb 100644 --- a/tests/example2.v +++ b/tests/example2.v @@ -13,3 +13,9 @@ Proof. split with 0. split. Qed. + +Goal exists n, n = 0. +Proof. +let myvar := Std.NamedHyp @x in split with (?myvar := 0). +split. +Qed. -- cgit v1.2.3 From 33e2bfe7a5eb9867634be82262ad041460709bcb Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 1 Aug 2017 20:52:52 +0200 Subject: Expanding unification variables in typechecking error messages. --- src/tac2intern.ml | 37 +++++++++++++++++++++++++++---------- 1 file changed, 27 insertions(+), 10 deletions(-) diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 16e0bc8cbe..32ed211ad0 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -316,6 +316,27 @@ let rec kind env t = match t with if is_unfoldable kn then kind env (unfold env kn tl) else t | GTypArrow _ | GTypRef (Tuple _, _) -> t +(** Normalize unification variables without unfolding type aliases *) +let rec nf env t = match t with +| GTypVar id -> + let (id, v) = UF.find id env.env_cst in + begin match v with + | None -> GTypVar id + | Some t -> nf env t + end +| GTypRef (kn, tl) -> + let tl = List.map (fun t -> nf env t) tl in + GTypRef (kn, tl) +| GTypArrow (t, u) -> + let t = nf env t in + let u = nf env u in + GTypArrow (t, u) + +let pr_glbtype env t = + let t = nf env t in + let name = env_name env in + pr_glbtype name t + exception Occur let rec occur_check env id t = match kind env t with @@ -357,9 +378,8 @@ let rec unify env t1 t2 = match kind env t1, kind env t2 with let unify ?loc env t1 t2 = try unify env t1 t2 with CannotUnify (u1, u2) -> - let name = env_name env in - user_err ?loc (str "This expression has type " ++ pr_glbtype name t1 ++ - str " but an expression was expected of type " ++ pr_glbtype name t2) + user_err ?loc (str "This expression has type " ++ pr_glbtype env t1 ++ + str " but an expression was expected of type " ++ pr_glbtype env t2) let unify_arrow ?loc env ft args = let ft0 = ft in @@ -373,12 +393,11 @@ let unify_arrow ?loc env ft args = let () = unify_var env id (GTypArrow (t, ft)) in iter ft args true | GTypRef _, _ :: _ -> - let name = env_name env in if is_fun then - user_err ?loc (str "This function has type " ++ pr_glbtype name ft0 ++ + user_err ?loc (str "This function has type " ++ pr_glbtype env ft0 ++ str " and is applied to too many arguments") else - user_err ?loc (str "This expression has type " ++ pr_glbtype name ft0 ++ + user_err ?loc (str "This expression has type " ++ pr_glbtype env ft0 ++ str " and is not a function") in iter ft args false @@ -478,15 +497,13 @@ let check_elt_empty loc env t = match kind env t with | GTypVar _ -> user_err ~loc (str "Cannot infer an empty type for this expression") | GTypArrow _ | GTypRef (Tuple _, _) -> - let name = env_name env in - user_err ~loc (str "Type " ++ pr_glbtype name t ++ str " is not an empty type") + user_err ~loc (str "Type " ++ pr_glbtype env t ++ str " is not an empty type") | GTypRef (Other kn, _) -> let def = Tac2env.interp_type kn in match def with | _, GTydAlg { galg_constructors = [] } -> kn | _ -> - let name = env_name env in - user_err ~loc (str "Type " ++ pr_glbtype name t ++ str " is not an empty type") + user_err ~loc (str "Type " ++ pr_glbtype env t ++ str " is not an empty type") let check_unit ?loc t = let env = empty_env () in -- cgit v1.2.3 From dd1343eb2680c202cf059e3db5788904b7d79782 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 1 Aug 2017 21:52:15 +0200 Subject: More primitive tactics. --- src/tac2stdlib.ml | 16 ++++++++++++++++ theories/Std.v | 3 +++ 2 files changed, 19 insertions(+) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 44fad48955..8fdf9c6d8c 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -196,6 +196,22 @@ let () = define_prim1 "tac_egeneralize" begin fun cl -> Tactics.new_generalize_gen cl end +let () = define_prim3 "tac_assert" begin fun c tac ipat -> + let c = Value.to_constr c in + let of_tac t = Proofview.tclIGNORE (thaw t) in + let tac = Value.to_option (fun t -> Value.to_option of_tac t) tac in + let ipat = Value.to_option (fun ipat -> Loc.tag (to_intro_pattern ipat)) ipat in + Tactics.forward true tac ipat c +end + +let () = define_prim3 "tac_enough" begin fun c tac ipat -> + let c = Value.to_constr c in + let of_tac t = Proofview.tclIGNORE (thaw t) in + let tac = Value.to_option (fun t -> Value.to_option of_tac t) tac in + let ipat = Value.to_option (fun ipat -> Loc.tag (to_intro_pattern ipat)) ipat in + Tactics.forward false tac ipat c +end + let () = define_prim2 "tac_pose" begin fun idopt c -> let na = to_name idopt in let c = Value.to_constr c in diff --git a/theories/Std.v b/theories/Std.v index a27790c35d..d2b85f215e 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -83,6 +83,9 @@ Ltac2 @ external ecase : constr_with_bindings -> unit := "ltac2" "tac_ecase". Ltac2 @ external egeneralize : (constr * occurrences * ident option) list -> unit := "ltac2" "tac_egeneralize". +Ltac2 @ external assert : constr -> (unit -> unit) option option -> intro_pattern option -> unit := "ltac2" "tac_assert". +Ltac2 @ external enough : constr -> (unit -> unit) option option -> intro_pattern option -> unit := "ltac2" "tac_enough". + Ltac2 @ external pose : ident option -> constr -> unit := "ltac2" "tac_pose". Ltac2 @ external set : ident option -> (unit -> constr) -> clause -> unit := "ltac2" "tac_set". Ltac2 @ external eset : ident option -> (unit -> constr) -> clause -> unit := "ltac2" "tac_eset". -- cgit v1.2.3 From a5419f01eb48b1cb3f5dee5482263530ad075ef4 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 2 Aug 2017 01:15:55 +0200 Subject: Fixup reification of egeneralize. --- src/tac2stdlib.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 8fdf9c6d8c..5c8337d41a 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -189,7 +189,7 @@ end let () = define_prim1 "tac_egeneralize" begin fun cl -> let cast = function | ValBlk (0, [| c; occs; na |]) -> - ((to_occurrences Value.to_int c, Value.to_constr c), to_name na) + ((to_occurrences Value.to_int occs, Value.to_constr c), to_name na) | _ -> assert false in let cl = Value.to_list cast cl in -- cgit v1.2.3 From b760af386d3c69c6963231489094685ea2a1e673 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 2 Aug 2017 00:43:20 +0200 Subject: Tentatively fixing a few parsing issues. --- src/g_ltac2.ml4 | 24 ++++++++++++++---------- 1 file changed, 14 insertions(+), 10 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 71fb59acf8..13d5dba8c6 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -83,29 +83,33 @@ GEXTEND Gram ] ; tac2expr: - [ "5" - [ "fun"; it = LIST1 input_fun ; "=>"; body = tac2expr LEVEL "5" -> CTacFun (!@loc, it, body) + [ "top" RIGHTA + [ e1 = SELF; ";"; e2 = SELF -> CTacSeq (!@loc, e1, e2) ] + | "5" + [ "fun"; it = LIST1 input_fun ; "=>"; body = tac2expr LEVEL "top" -> CTacFun (!@loc, it, body) | "let"; isrec = rec_flag; lc = LIST1 let_clause SEP "with"; "in"; - e = tac2expr LEVEL "5" -> CTacLet (!@loc, isrec, lc, e) + e = tac2expr LEVEL "top" -> CTacLet (!@loc, isrec, lc, e) | "match"; e = tac2expr LEVEL "5"; "with"; bl = branches ;"end" -> CTacCse (!@loc, e, bl) ] - | "2" LEFTA - [ e1 = tac2expr; ";"; e2 = tac2expr -> CTacSeq (!@loc, e1, e2) ] + | "::" RIGHTA + [ e1 = tac2expr; "::"; e2 = tac2expr -> + CTacApp (!@loc, CTacCst (!@loc, AbsKn (Other Tac2core.Core.c_cons)), [e1; e2]) + ] + | [ e0 = SELF; ","; el = LIST1 NEXT SEP "," -> + let el = e0 :: el in + CTacApp (!@loc, CTacCst (!@loc, AbsKn (Tuple (List.length el))), el) ] | "1" LEFTA [ e = tac2expr; el = LIST1 tac2expr LEVEL "0" -> CTacApp (!@loc, e, el) | e = SELF; ".("; qid = Prim.qualid; ")" -> CTacPrj (!@loc, e, RelId qid) - | e = SELF; ".("; qid = Prim.qualid; ")"; ":="; r = tac2expr LEVEL "1" -> CTacSet (!@loc, e, RelId qid, r) - | e0 = tac2expr; ","; el = LIST1 tac2expr LEVEL "1" SEP "," -> - let el = e0 :: el in - CTacApp (!@loc, CTacCst (!@loc, AbsKn (Tuple (List.length el))), el) ] + | e = SELF; ".("; qid = Prim.qualid; ")"; ":="; r = tac2expr LEVEL "5" -> CTacSet (!@loc, e, RelId qid, r) ] | "0" [ "("; a = tac2expr LEVEL "5"; ")" -> a | "("; a = tac2expr; ":"; t = tac2type; ")" -> CTacCnv (!@loc, a, t) | "()" -> CTacCst (!@loc, AbsKn (Tuple 0)) | "("; ")" -> CTacCst (!@loc, AbsKn (Tuple 0)) - | "["; a = LIST0 tac2expr LEVEL "1" SEP ";"; "]" -> CTacLst (Loc.tag ~loc:!@loc a) + | "["; a = LIST0 tac2expr LEVEL "5" SEP ";"; "]" -> CTacLst (Loc.tag ~loc:!@loc a) | "{"; a = tac2rec_fieldexprs; "}" -> CTacRec (!@loc, a) | a = tactic_atom -> a ] ] -- cgit v1.2.3 From d3c3859ab6dba6495b13e055917ddf3d95851912 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 2 Aug 2017 01:26:17 +0200 Subject: Better test Makefile. --- tests/Makefile | 7 ++++++- 1 file changed, 6 insertions(+), 1 deletion(-) diff --git a/tests/Makefile b/tests/Makefile index a48ab0860f..9370b063f8 100644 --- a/tests/Makefile +++ b/tests/Makefile @@ -1,7 +1,12 @@ all: $(patsubst %.v,%.v.log,$(wildcard *.v)) %.v.log: %.v - $(COQBIN)/coqtop -I ../src -Q ../theories Ltac2 < $< 2> $@ + $(COQBIN)/coqtop -batch -I ../src -Q ../theories Ltac2 -lv $< > $@ + if [ $$? = 0 ]; then \ + echo " $<... OK"; \ + else \ + echo " $<... FAIL!"; \ + fi; \ clean: rm -f *.log -- cgit v1.2.3 From 087012f8d3e5e31f489e35dce8397b5202c928b6 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 2 Aug 2017 01:57:48 +0200 Subject: Adding the open_constr scope --- src/tac2core.ml | 1 + src/tac2quote.ml | 2 ++ src/tac2quote.mli | 2 ++ 3 files changed, 5 insertions(+) diff --git a/src/tac2core.ml b/src/tac2core.ml index 266e3b5f11..d2cc865299 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -876,3 +876,4 @@ let () = add_scope "intropatterns" begin function end let () = add_generic_scope "constr" Pcoq.Constr.constr Stdarg.wit_constr +let () = add_generic_scope "open_constr" Pcoq.Constr.constr Stdarg.wit_open_constr diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 488bcb5201..0e0a7b3fce 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -55,6 +55,8 @@ let of_ident ?loc id = inj_wit ?loc Stdarg.wit_ident id let of_constr ?loc c = inj_wit ?loc Stdarg.wit_constr c +let of_open_constr ?loc c = inj_wit ?loc Stdarg.wit_open_constr c + let of_bool ?loc b = let c = if b then Core.c_true else Core.c_false in constructor ?loc c [] diff --git a/src/tac2quote.mli b/src/tac2quote.mli index c9ee270d38..a311430a66 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -30,6 +30,8 @@ val of_ident : ?loc:Loc.t -> Id.t -> raw_tacexpr val of_constr : ?loc:Loc.t -> Constrexpr.constr_expr -> raw_tacexpr +val of_open_constr : ?loc:Loc.t -> Constrexpr.constr_expr -> raw_tacexpr + val of_list : ?loc:Loc.t -> raw_tacexpr list -> raw_tacexpr val of_bindings : ?loc:Loc.t -> bindings -> raw_tacexpr -- cgit v1.2.3 From c96b746b17a37e242fc01103d22fa0b852da84c5 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 2 Aug 2017 01:59:32 +0200 Subject: Bindings use open constr quotations. --- src/g_ltac2.ml4 | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 13d5dba8c6..3c41828cd3 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -301,18 +301,18 @@ GEXTEND Gram ; simple_binding: [ [ "("; LEFTQMARK; id = Prim.ident; ":="; c = Constr.lconstr; ")" -> - Loc.tag ~loc:!@loc (QAnti (Loc.tag ~loc:!@loc id), Tac2quote.of_constr ~loc:!@loc c) + Loc.tag ~loc:!@loc (QAnti (Loc.tag ~loc:!@loc id), Tac2quote.of_open_constr ~loc:!@loc c) | "("; n = Prim.natural; ":="; c = Constr.lconstr; ")" -> - Loc.tag ~loc:!@loc (QExpr (AnonHyp n), Tac2quote.of_constr ~loc:!@loc c) + Loc.tag ~loc:!@loc (QExpr (AnonHyp n), Tac2quote.of_open_constr ~loc:!@loc c) | "("; id = Prim.ident; ":="; c = Constr.lconstr; ")" -> - Loc.tag ~loc:!@loc (QExpr (NamedHyp id), Tac2quote.of_constr ~loc:!@loc c) + Loc.tag ~loc:!@loc (QExpr (NamedHyp id), Tac2quote.of_open_constr ~loc:!@loc c) ] ] ; bindings: [ [ test_lpar_idnum_coloneq; bl = LIST1 simple_binding -> QExplicitBindings bl | bl = LIST1 Constr.constr -> - let bl = List.map (fun c -> Tac2quote.of_constr ~loc:!@loc c) bl in + let bl = List.map (fun c -> Tac2quote.of_open_constr ~loc:!@loc c) bl in QImplicitBindings bl ] ] ; -- cgit v1.2.3 From 53374f189cc9b9b67ff94d5362fdffdba6c779a3 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 2 Aug 2017 01:42:14 +0200 Subject: Adding a few standard notations for Ltac1 tactics. --- _CoqProject | 1 + tests/example2.v | 2 +- theories/Ltac2.v | 1 + theories/Notations.v | 37 +++++++++++++++++++++++++++++++++++++ 4 files changed, 40 insertions(+), 1 deletion(-) create mode 100644 theories/Notations.v diff --git a/_CoqProject b/_CoqProject index b8064c46a4..583639612b 100644 --- a/_CoqProject +++ b/_CoqProject @@ -35,4 +35,5 @@ theories/Message.v theories/Constr.v theories/Pattern.v theories/Std.v +theories/Notations.v theories/Ltac2.v diff --git a/tests/example2.v b/tests/example2.v index ffdb723ffb..398f33561e 100644 --- a/tests/example2.v +++ b/tests/example2.v @@ -1,6 +1,6 @@ Require Import Ltac2.Ltac2. -Ltac2 Notation "split" bnd(bindings) := Std.split bnd. +Import Ltac2.Notations. Goal exists n, n = 0. Proof. diff --git a/theories/Ltac2.v b/theories/Ltac2.v index 9aaee850cd..07229797da 100644 --- a/theories/Ltac2.v +++ b/theories/Ltac2.v @@ -17,3 +17,4 @@ Require Ltac2.Constr. Require Ltac2.Control. Require Ltac2.Pattern. Require Ltac2.Std. +Require Ltac2.Notations. diff --git a/theories/Notations.v b/theories/Notations.v new file mode 100644 index 0000000000..d0400667db --- /dev/null +++ b/theories/Notations.v @@ -0,0 +1,37 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Std.split bnd). + +Ltac2 Notation "esplit" bnd(bindings) := Std.esplit bnd. + +Ltac2 Notation "left" bnd(thunk(bindings)) := + Control.with_holes bnd (fun bnd => Std.left bnd). + +Ltac2 Notation "eleft" bnd(bindings) := Std.eleft bnd. + +Ltac2 Notation "right" bnd(thunk(bindings)) := + Control.with_holes bnd (fun bnd => Std.right bnd). + +Ltac2 Notation "eright" bnd(bindings) := Std.eright bnd. + +Ltac2 Notation "constructor" := Std.constructor (). +Ltac2 Notation "constructor" n(tactic) bnd(thunk(bindings)) := + Control.with_holes bnd (fun bnd => Std.constructor_n n bnd). + +Ltac2 Notation "econstructor" := Std.econstructor (). +Ltac2 Notation "econstructor" n(tactic) bnd(bindings) := + Std.econstructor_n n bnd. -- cgit v1.2.3 From d0766f4dc08b00128a47a00ca74334ba0bfeed24 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 2 Aug 2017 12:15:05 +0200 Subject: Removing deprecated stuff. --- src/tac2print.mli | 14 +++++++------- 1 file changed, 7 insertions(+), 7 deletions(-) diff --git a/src/tac2print.mli b/src/tac2print.mli index ddd599641d..2ee5cf42e0 100644 --- a/src/tac2print.mli +++ b/src/tac2print.mli @@ -19,16 +19,16 @@ type typ_level = | T1 | T0 -val pr_typref : type_constant -> std_ppcmds -val pr_glbtype_gen : ('a -> string) -> typ_level -> 'a glb_typexpr -> std_ppcmds -val pr_glbtype : ('a -> string) -> 'a glb_typexpr -> std_ppcmds +val pr_typref : type_constant -> Pp.t +val pr_glbtype_gen : ('a -> string) -> typ_level -> 'a glb_typexpr -> Pp.t +val pr_glbtype : ('a -> string) -> 'a glb_typexpr -> Pp.t (** {5 Printing expressions} *) -val pr_constructor : ltac_constructor -> std_ppcmds -val pr_projection : ltac_projection -> std_ppcmds -val pr_glbexpr_gen : exp_level -> glb_tacexpr -> std_ppcmds -val pr_glbexpr : glb_tacexpr -> std_ppcmds +val pr_constructor : ltac_constructor -> Pp.t +val pr_projection : ltac_projection -> Pp.t +val pr_glbexpr_gen : exp_level -> glb_tacexpr -> Pp.t +val pr_glbexpr : glb_tacexpr -> Pp.t (** {5 Utilities} *) -- cgit v1.2.3 From ebee89f2b2d1815dbb89916363de1b1ad17890e8 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 2 Aug 2017 13:25:12 +0200 Subject: Fixing parsing of match branches. --- src/g_ltac2.ml4 | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 3c41828cd3..20a00afa2e 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -90,7 +90,7 @@ GEXTEND Gram | "let"; isrec = rec_flag; lc = LIST1 let_clause SEP "with"; "in"; e = tac2expr LEVEL "top" -> CTacLet (!@loc, isrec, lc, e) - | "match"; e = tac2expr LEVEL "5"; "with"; bl = branches ;"end" -> + | "match"; e = tac2expr LEVEL "5"; "with"; bl = branches; "end" -> CTacCse (!@loc, e, bl) ] | "::" RIGHTA @@ -121,7 +121,7 @@ GEXTEND Gram ] ; branch: - [ [ pat = tac2pat LEVEL "1"; "=>"; e = tac2expr LEVEL "5" -> (pat, e) ] ] + [ [ pat = tac2pat LEVEL "1"; "=>"; e = tac2expr LEVEL "top" -> (pat, e) ] ] ; rec_flag: [ [ IDENT "rec" -> true -- cgit v1.2.3 From e50d86c836cf492a637db056b446bb4c70b2e40b Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 2 Aug 2017 13:29:12 +0200 Subject: Properly classifying Ltac2 notations. --- src/g_ltac2.ml4 | 9 ++++++--- 1 file changed, 6 insertions(+), 3 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 20a00afa2e..f558f9b9cc 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -397,8 +397,12 @@ PRINTED BY pr_ltac2entry | [ tac2def_syn(e) ] -> [ e ] END -VERNAC COMMAND EXTEND VernacDeclareTactic2Definition CLASSIFIED AS SIDEFF -| [ "Ltac2" ltac2_entry(e) ] -> [ +let classify_ltac2 = function +| StrSyn _ -> Vernacexpr.VtUnknown, Vernacexpr.VtNow +| StrVal _ | StrPrm _ | StrTyp _ -> Vernac_classifier.classify_as_sideeff + +VERNAC COMMAND EXTEND VernacDeclareTactic2Definition +| [ "Ltac2" ltac2_entry(e) ] => [ classify_ltac2 e ] -> [ let local = Locality.LocalityFixme.consume () in Tac2entries.register_struct ?local e ] @@ -433,4 +437,3 @@ open Stdarg VERNAC COMMAND EXTEND Ltac2Print CLASSIFIED AS SIDEFF | [ "Print" "Ltac2" reference(tac) ] -> [ Tac2entries.print_ltac tac ] END - -- cgit v1.2.3 From da8eec98d095482c0e12c0ece9725a300e5f3d57 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 2 Aug 2017 13:49:48 +0200 Subject: More examples --- tests/stuff/ltac2.v | 13 +------------ tests/typing.v | 47 +++++++++++++++++++++++++++++++++++++++++++++++ 2 files changed, 48 insertions(+), 12 deletions(-) diff --git a/tests/stuff/ltac2.v b/tests/stuff/ltac2.v index 6b30d42c09..95bff7e569 100644 --- a/tests/stuff/ltac2.v +++ b/tests/stuff/ltac2.v @@ -55,10 +55,6 @@ Ltac2 qux3 x := constr:(nat -> ltac2:(refine (fun () => hyp x))). Print Ltac2 qux3. -Ltac2 qux4 f x := x, (f x, x). - -Print Ltac2 qux4. - Ltac2 Type rec nat := [ O | S (nat) ]. Ltac2 message_of_nat n := @@ -137,17 +133,10 @@ Std.ecase (hyp @H, Std.ExplicitBindings [Std.NamedHyp @n, '0]). refine (fun () => 'eq_refl). Qed. -Ltac2 rec do n tac := match Int.equal n 0 with -| true => () -| false => tac (); do (Int.sub n 1) tac -end. - -Print Ltac2 do. - Goal forall x, 1 + x = x + 1. Proof. refine (fun () => '(fun x => _)). Std.cbv { Std.rBeta := true; Std.rMatch := true; Std.rFix := true; Std.rCofix := true; - Std.rZeta := true; Std.rDelta := false; rConst := []; + Std.rZeta := true; Std.rDelta := false; Std.rConst := []; } { Std.on_hyps := None; Std.on_concl := Std.AllOccurrences }. diff --git a/tests/typing.v b/tests/typing.v index 8460ab42b7..9f18292716 100644 --- a/tests/typing.v +++ b/tests/typing.v @@ -14,6 +14,10 @@ Fail Ltac2 test2 () := test0 true. Fail Ltac2 test2 () := test0 0 0. +Ltac2 test3 f x := x, (f x, x). + +Print Ltac2 test3. + (** Polymorphism *) Ltac2 rec list_length l := @@ -23,3 +27,46 @@ match l with end. Print Ltac2 list_length. + +(** Pattern-matching *) + +Ltac2 ifb b f g := match b with +| true => f () +| false => g () +end. + +Print Ltac2 ifb. + +Ltac2 if_not_found e f g := match e with +| Not_found => f () +| _ => g () +end. + +Fail Ltac2 ifb' b f g := match b with +| true => f () +end. + +Fail Ltac2 if_not_found' e f g := match e with +| Not_found => f () +end. + +(** Reimplementing 'do'. Return value of the function useless. *) + +Ltac2 rec do n tac := match Int.equal n 0 with +| true => () +| false => tac (); do (Int.sub n 1) tac +end. + +Print Ltac2 do. + +(** Non-function pure values are OK. *) + +Ltac2 tuple0 := ([1; 2], true, (fun () => "yay")). + +Print Ltac2 tuple0. + +(** Impure values are not. *) + +Fail Ltac2 not_a_value := { contents := 0 }. +Fail Ltac2 not_a_value := "nope". +Fail Ltac2 not_a_value := list_length []. -- cgit v1.2.3 From ea782d757d57dc31be9714edc607128c5c127205 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 2 Aug 2017 14:14:14 +0200 Subject: Extending the set of tactic scopes. We now allow mere tokens, keywords and sequencing amongst others. --- doc/ltac2.md | 13 ++++++++- src/tac2core.ml | 85 ++++++++++++++++++++++++++++++++++++++++++++++++++++++ src/tac2entries.ml | 3 ++ 3 files changed, 100 insertions(+), 1 deletion(-) diff --git a/doc/ltac2.md b/doc/ltac2.md index 7a3b2181f8..12687e9aff 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -552,8 +552,19 @@ The following scopes are built-in. - tactic(n = INT): + parses a Ltac2 expression at the provided level *n* and return it as is. - thunk(*scope*): - parses the same as *scope*, and if *e* is the parsed expression, returns + + parses the same as *scope*, and if *e* is the parsed expression, returns `fun () => e`. +- STRING: + + parses the corresponding string as a CAMLP5 IDENT and returns `()`. +- keyword(s = STRING): + + parses the string *s* as a keyword and returns `()`. +- terminal(s = STRING): + + parses the string *s* as a keyword, if it is already a + keyword, otherwise as an IDENT. Returns `()`. +- seq(*scope₁*, ..., *scopeₙ*): + + parses *scope₁*, ..., *scopeₙ* in this order, and produces a n-tuple made + out of the parsed values in the same order. It is forbidden for the various + subscopes to refer to the global entry using self of next. For now there is no way to declare new scopes from Ltac2 side, but this is planned. diff --git a/src/tac2core.ml b/src/tac2core.ml index d2cc865299..b45275210e 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -742,6 +742,8 @@ let scope_fail () = CErrors.user_err (str "Invalid parsing token") let dummy_loc = Loc.make_loc (-1, -1) +let q_unit = CTacCst (dummy_loc, AbsKn (Tuple 0)) + let rthunk e = let loc = Tac2intern.loc_of_tacexpr e in let var = [CPatVar (Some loc, Anonymous), Some (CTypRef (loc, AbsKn (Other Core.t_unit), []))] in @@ -757,6 +759,20 @@ let add_generic_scope s entry arg = in add_scope s parse +let () = add_scope "keyword" begin function +| [SexprStr (loc, s)] -> + let scope = Extend.Atoken (Tok.KEYWORD s) in + Tac2entries.ScopeRule (scope, (fun _ -> q_unit)) +| _ -> scope_fail () +end + +let () = add_scope "terminal" begin function +| [SexprStr (loc, s)] -> + let scope = Extend.Atoken (CLexer.terminal s) in + Tac2entries.ScopeRule (scope, (fun _ -> q_unit)) +| _ -> scope_fail () +end + let () = add_scope "list0" begin function | [tok] -> let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in @@ -877,3 +893,72 @@ end let () = add_generic_scope "constr" Pcoq.Constr.constr Stdarg.wit_constr let () = add_generic_scope "open_constr" Pcoq.Constr.constr Stdarg.wit_open_constr + +(** seq scope, a bit hairy *) + +open Extend +exception SelfSymbol + +type 'a any_symbol = { any_symbol : 'r. ('r, 'a) symbol } + +let rec generalize_symbol : + type a s. (s, a) Extend.symbol -> a any_symbol = function +| Atoken tok -> + { any_symbol = Atoken tok } +| Alist1 e -> + let e = generalize_symbol e in + { any_symbol = Alist1 e.any_symbol } +| Alist1sep (e, sep) -> + let e = generalize_symbol e in + let sep = generalize_symbol sep in + { any_symbol = Alist1sep (e.any_symbol, sep.any_symbol) } +| Alist0 e -> + let e = generalize_symbol e in + { any_symbol = Alist0 e.any_symbol } +| Alist0sep (e, sep) -> + let e = generalize_symbol e in + let sep = generalize_symbol sep in + { any_symbol = Alist0sep (e.any_symbol, sep.any_symbol) } +| Aopt e -> + let e = generalize_symbol e in + { any_symbol = Aopt e.any_symbol } +| Aself -> raise SelfSymbol +| Anext -> raise SelfSymbol +| Aentry e -> { any_symbol = Aentry e } +| Aentryl (e, l) -> { any_symbol = Aentryl (e, l) } +| Arules r -> { any_symbol = Arules r } + +type _ converter = +| CvNil : (Loc.t -> raw_tacexpr) converter +| CvCns : 'act converter * ('a -> raw_tacexpr) -> ('a -> 'act) converter + +let rec apply : type a. a converter -> raw_tacexpr list -> a = function +| CvNil -> fun accu loc -> + let cst = CTacCst (loc, AbsKn (Tuple (List.length accu))) in + CTacApp (loc, cst, accu) +| CvCns (c, f) -> fun accu x -> apply c (f x :: accu) + +type seqrule = +| Seqrule : ('act, Loc.t -> raw_tacexpr) norec_rule * 'act converter -> seqrule + +let rec make_seq_rule = function +| [] -> + let r = { norec_rule = Stop } in + Seqrule (r, CvNil) +| tok :: rem -> + let Tac2entries.ScopeRule (scope, f) = Tac2entries.parse_scope tok in + let scope = generalize_symbol scope in + let Seqrule (r, c) = make_seq_rule rem in + let r = { norec_rule = Next (r.norec_rule, scope.any_symbol) } in + Seqrule (r, CvCns (c, f)) + +let () = add_scope "seq" begin fun toks -> + let scope = + try + let Seqrule (r, c) = make_seq_rule (List.rev toks) in + Arules [Rules (r, apply c [])] + with SelfSymbol -> + CErrors.user_err (str "Recursive symbols (self / next) are not allowed in local rules") + in + Tac2entries.ScopeRule (scope, (fun e -> e)) +end diff --git a/src/tac2entries.ml b/src/tac2entries.ml index d7ee07e9e2..0f32736096 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -507,6 +507,9 @@ let parse_scope = function Id.Map.find id !scope_table toks else CErrors.user_err ?loc (str "Unknown scope" ++ spc () ++ Nameops.pr_id id) +| SexprStr (_, str) -> + let v_unit = CTacCst (dummy_loc, AbsKn (Tuple 0)) in + ScopeRule (Extend.Atoken (Tok.IDENT str), (fun _ -> v_unit)) | tok -> let loc = loc_of_token tok in CErrors.user_err ~loc (str "Invalid parsing token") -- cgit v1.2.3 From faf40da077f20a67a45fe98f8ef99f90440ef16d Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 2 Aug 2017 16:32:10 +0200 Subject: Adding new notations. --- src/tac2stdlib.ml | 11 +++++++++++ theories/Notations.v | 26 ++++++++++++++++++++++++++ theories/Std.v | 2 ++ 3 files changed, 39 insertions(+) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 5c8337d41a..ac530f5130 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -175,12 +175,23 @@ let () = define_prim1 "tac_eintros" begin fun ipat -> Tactics.intros_patterns true ipat end +let () = define_prim2 "tac_elim" begin fun c copt -> + let c = to_constr_with_bindings c in + let copt = Value.to_option to_constr_with_bindings copt in + Tactics.elim false None c copt +end + let () = define_prim2 "tac_eelim" begin fun c copt -> let c = to_constr_with_bindings c in let copt = Value.to_option to_constr_with_bindings copt in Tactics.elim true None c copt end +let () = define_prim1 "tac_case" begin fun c -> + let c = to_constr_with_bindings c in + Tactics.general_case_analysis false None c +end + let () = define_prim1 "tac_ecase" begin fun c -> let c = to_constr_with_bindings c in Tactics.general_case_analysis true None c diff --git a/theories/Notations.v b/theories/Notations.v index d0400667db..ec7a6b0b12 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -35,3 +35,29 @@ Ltac2 Notation "constructor" n(tactic) bnd(thunk(bindings)) := Ltac2 Notation "econstructor" := Std.econstructor (). Ltac2 Notation "econstructor" n(tactic) bnd(bindings) := Std.econstructor_n n bnd. + +Ltac2 eelim c bnd use := + let use := match use with + | None => None + | Some u => + let ((_, c, wth)) := u in Some (c, wth) + end in + Std.eelim (c, bnd) use. + +Ltac2 elim c bnd use := + Control.with_holes + (fun () => c (), bnd (), use ()) + (fun ((c, bnd, use)) => + let use := match use with + | None => None + | Some u => + let ((_, c, wth)) := u in Some (c, wth) + end in + Std.elim (c, bnd) use). + +Ltac2 Notation "elim" c(thunk(constr)) bnd(thunk(bindings)) + use(thunk(opt(seq("using", constr, bindings)))) := elim c bnd use. + +Ltac2 Notation "eelim" c(constr) bnd(bindings) + use(opt(seq("using", constr, bindings))) := + eelim c bnd use. diff --git a/theories/Std.v b/theories/Std.v index d2b85f215e..3d0f463c5e 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -78,7 +78,9 @@ with or_and_intro_pattern := [ Ltac2 @ external intros : intro_pattern list -> unit := "ltac2" "tac_intros". Ltac2 @ external eintros : intro_pattern list -> unit := "ltac2" "tac_eintros". +Ltac2 @ external elim : constr_with_bindings -> constr_with_bindings option -> unit := "ltac2" "tac_elim". Ltac2 @ external eelim : constr_with_bindings -> constr_with_bindings option -> unit := "ltac2" "tac_eelim". +Ltac2 @ external case : constr_with_bindings -> unit := "ltac2" "tac_case". Ltac2 @ external ecase : constr_with_bindings -> unit := "ltac2" "tac_ecase". Ltac2 @ external egeneralize : (constr * occurrences * ident option) list -> unit := "ltac2" "tac_egeneralize". -- cgit v1.2.3 From 9088f6db4f56d906d8a18eeaf09c9adbae4a5fd4 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 2 Aug 2017 16:46:12 +0200 Subject: Merging the e/- variants of primitive tactics. --- src/tac2stdlib.ml | 97 +++++++++++++++++++--------------------------------- theories/Notations.v | 28 +++++++-------- theories/Std.v | 31 +++++++---------- 3 files changed, 61 insertions(+), 95 deletions(-) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index ac530f5130..f63252ec22 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -163,41 +163,35 @@ let define_prim3 name tac = in Tac2env.define_primitive (pname name) tac -(** Tactics from Tacexpr *) +let define_prim4 name tac = + let tac = function + | [x; y; z; u] -> lift (tac x y z u) + | _ -> assert false + in + Tac2env.define_primitive (pname name) tac -let () = define_prim1 "tac_intros" begin fun ipat -> - let ipat = to_intro_patterns ipat in - Tactics.intros_patterns false ipat -end +(** Tactics from Tacexpr *) -let () = define_prim1 "tac_eintros" begin fun ipat -> +let () = define_prim2 "tac_intros" begin fun ev ipat -> + let ev = Value.to_bool ev in let ipat = to_intro_patterns ipat in - Tactics.intros_patterns true ipat + Tactics.intros_patterns ev ipat end -let () = define_prim2 "tac_elim" begin fun c copt -> +let () = define_prim3 "tac_elim" begin fun ev c copt -> + let ev = Value.to_bool ev in let c = to_constr_with_bindings c in let copt = Value.to_option to_constr_with_bindings copt in - Tactics.elim false None c copt + Tactics.elim ev None c copt end -let () = define_prim2 "tac_eelim" begin fun c copt -> +let () = define_prim2 "tac_case" begin fun ev c -> + let ev = Value.to_bool ev in let c = to_constr_with_bindings c in - let copt = Value.to_option to_constr_with_bindings copt in - Tactics.elim true None c copt -end - -let () = define_prim1 "tac_case" begin fun c -> - let c = to_constr_with_bindings c in - Tactics.general_case_analysis false None c + Tactics.general_case_analysis ev None c end -let () = define_prim1 "tac_ecase" begin fun c -> - let c = to_constr_with_bindings c in - Tactics.general_case_analysis true None c -end - -let () = define_prim1 "tac_egeneralize" begin fun cl -> +let () = define_prim1 "tac_generalize" begin fun cl -> let cast = function | ValBlk (0, [| c; occs; na |]) -> ((to_occurrences Value.to_int occs, Value.to_constr c), to_name na) @@ -229,22 +223,14 @@ let () = define_prim2 "tac_pose" begin fun idopt c -> Tactics.letin_tac None na c None Locusops.nowhere end -let () = define_prim3 "tac_set" begin fun idopt c cl -> - let na = to_name idopt in - let cl = to_clause cl in - Proofview.tclEVARMAP >>= fun sigma -> - thaw c >>= fun c -> - let c = Value.to_constr c in - Tactics.letin_pat_tac false None na (sigma, c) cl -end - -let () = define_prim3 "tac_eset" begin fun idopt c cl -> +let () = define_prim4 "tac_set" begin fun ev idopt c cl -> + let ev = Value.to_bool ev in let na = to_name idopt in let cl = to_clause cl in Proofview.tclEVARMAP >>= fun sigma -> thaw c >>= fun c -> let c = Value.to_constr c in - Tactics.letin_pat_tac true None na (sigma, c) cl + Tactics.letin_pat_tac ev None na (sigma, c) cl end let () = define_prim1 "tac_red" begin fun cl -> @@ -301,21 +287,15 @@ let () = define_prim1 "tac_cut" begin fun c -> Tactics.cut c end -let () = define_prim1 "tac_left" begin fun bnd -> - let bnd = to_bindings bnd in - Tactics.left_with_bindings false bnd -end -let () = define_prim1 "tac_eleft" begin fun bnd -> +let () = define_prim2 "tac_left" begin fun ev bnd -> + let ev = Value.to_bool ev in let bnd = to_bindings bnd in - Tactics.left_with_bindings true bnd + Tactics.left_with_bindings ev bnd end -let () = define_prim1 "tac_right" begin fun bnd -> +let () = define_prim2 "tac_right" begin fun ev bnd -> + let ev = Value.to_bool ev in let bnd = to_bindings bnd in - Tactics.right_with_bindings false bnd -end -let () = define_prim1 "tac_eright" begin fun bnd -> - let bnd = to_bindings bnd in - Tactics.right_with_bindings true bnd + Tactics.right_with_bindings ev bnd end let () = define_prim1 "tac_introsuntil" begin fun h -> @@ -334,19 +314,16 @@ let () = define_prim1 "tac_nativecastnocheck" begin fun c -> Tactics.native_cast_no_check (Value.to_constr c) end -let () = define_prim0 "tac_constructor" (Tactics.any_constructor false None) -let () = define_prim0 "tac_econstructor" (Tactics.any_constructor true None) - -let () = define_prim2 "tac_constructorn" begin fun n bnd -> - let n = Value.to_int n in - let bnd = to_bindings bnd in - Tactics.constructor_tac false None n bnd +let () = define_prim1 "tac_constructor" begin fun ev -> + let ev = Value.to_bool ev in + Tactics.any_constructor ev None end -let () = define_prim2 "tac_econstructorn" begin fun n bnd -> +let () = define_prim3 "tac_constructorn" begin fun ev n bnd -> + let ev = Value.to_bool ev in let n = Value.to_int n in let bnd = to_bindings bnd in - Tactics.constructor_tac true None n bnd + Tactics.constructor_tac ev None n bnd end let () = define_prim1 "tac_symmetry" begin fun cl -> @@ -354,14 +331,10 @@ let () = define_prim1 "tac_symmetry" begin fun cl -> Tactics.intros_symmetry cl end -let () = define_prim1 "tac_split" begin fun bnd -> - let bnd = to_bindings bnd in - Tactics.split_with_bindings false [bnd] -end - -let () = define_prim1 "tac_esplit" begin fun bnd -> +let () = define_prim2 "tac_split" begin fun ev bnd -> + let ev = Value.to_bool ev in let bnd = to_bindings bnd in - Tactics.split_with_bindings true [bnd] + Tactics.split_with_bindings ev [bnd] end let () = define_prim1 "tac_rename" begin fun ids -> diff --git a/theories/Notations.v b/theories/Notations.v index ec7a6b0b12..0487e324ca 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -9,32 +9,32 @@ Require Import Ltac2.Init. Require Ltac2.Control Ltac2.Std. -Ltac2 Notation "intros" p(intropatterns) := Std.intros p. +Ltac2 Notation "intros" p(intropatterns) := Std.intros false p. -Ltac2 Notation "eintros" p(intropatterns) := Std.eintros p. +Ltac2 Notation "eintros" p(intropatterns) := Std.intros true p. Ltac2 Notation "split" bnd(thunk(bindings)) := - Control.with_holes bnd (fun bnd => Std.split bnd). + Control.with_holes bnd (fun bnd => Std.split false bnd). -Ltac2 Notation "esplit" bnd(bindings) := Std.esplit bnd. +Ltac2 Notation "esplit" bnd(bindings) := Std.split true bnd. Ltac2 Notation "left" bnd(thunk(bindings)) := - Control.with_holes bnd (fun bnd => Std.left bnd). + Control.with_holes bnd (fun bnd => Std.left false bnd). -Ltac2 Notation "eleft" bnd(bindings) := Std.eleft bnd. +Ltac2 Notation "eleft" bnd(bindings) := Std.left true bnd. Ltac2 Notation "right" bnd(thunk(bindings)) := - Control.with_holes bnd (fun bnd => Std.right bnd). + Control.with_holes bnd (fun bnd => Std.right false bnd). -Ltac2 Notation "eright" bnd(bindings) := Std.eright bnd. +Ltac2 Notation "eright" bnd(bindings) := Std.right true bnd. -Ltac2 Notation "constructor" := Std.constructor (). +Ltac2 Notation "constructor" := Std.constructor false. Ltac2 Notation "constructor" n(tactic) bnd(thunk(bindings)) := - Control.with_holes bnd (fun bnd => Std.constructor_n n bnd). + Control.with_holes bnd (fun bnd => Std.constructor_n false n bnd). -Ltac2 Notation "econstructor" := Std.econstructor (). +Ltac2 Notation "econstructor" := Std.constructor true. Ltac2 Notation "econstructor" n(tactic) bnd(bindings) := - Std.econstructor_n n bnd. + Std.constructor_n true n bnd. Ltac2 eelim c bnd use := let use := match use with @@ -42,7 +42,7 @@ Ltac2 eelim c bnd use := | Some u => let ((_, c, wth)) := u in Some (c, wth) end in - Std.eelim (c, bnd) use. + Std.elim true (c, bnd) use. Ltac2 elim c bnd use := Control.with_holes @@ -53,7 +53,7 @@ Ltac2 elim c bnd use := | Some u => let ((_, c, wth)) := u in Some (c, wth) end in - Std.elim (c, bnd) use). + Std.elim false (c, bnd) use). Ltac2 Notation "elim" c(thunk(constr)) bnd(thunk(bindings)) use(thunk(opt(seq("using", constr, bindings)))) := elim c bnd use. diff --git a/theories/Std.v b/theories/Std.v index 3d0f463c5e..20504f1247 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -73,24 +73,22 @@ with or_and_intro_pattern := [ | IntroAndPattern (intro_pattern list) ]. +Ltac2 Type evar_flag := bool. + (** Standard, built-in tactics. See Ltac1 for documentation. *) -Ltac2 @ external intros : intro_pattern list -> unit := "ltac2" "tac_intros". -Ltac2 @ external eintros : intro_pattern list -> unit := "ltac2" "tac_eintros". +Ltac2 @ external intros : evar_flag -> intro_pattern list -> unit := "ltac2" "tac_intros". -Ltac2 @ external elim : constr_with_bindings -> constr_with_bindings option -> unit := "ltac2" "tac_elim". -Ltac2 @ external eelim : constr_with_bindings -> constr_with_bindings option -> unit := "ltac2" "tac_eelim". -Ltac2 @ external case : constr_with_bindings -> unit := "ltac2" "tac_case". -Ltac2 @ external ecase : constr_with_bindings -> unit := "ltac2" "tac_ecase". +Ltac2 @ external elim : evar_flag -> constr_with_bindings -> constr_with_bindings option -> unit := "ltac2" "tac_elim". +Ltac2 @ external case : evar_flag -> constr_with_bindings -> unit := "ltac2" "tac_case". -Ltac2 @ external egeneralize : (constr * occurrences * ident option) list -> unit := "ltac2" "tac_egeneralize". +Ltac2 @ external generalize : (constr * occurrences * ident option) list -> unit := "ltac2" "tac_generalize". Ltac2 @ external assert : constr -> (unit -> unit) option option -> intro_pattern option -> unit := "ltac2" "tac_assert". Ltac2 @ external enough : constr -> (unit -> unit) option option -> intro_pattern option -> unit := "ltac2" "tac_enough". Ltac2 @ external pose : ident option -> constr -> unit := "ltac2" "tac_pose". -Ltac2 @ external set : ident option -> (unit -> constr) -> clause -> unit := "ltac2" "tac_set". -Ltac2 @ external eset : ident option -> (unit -> constr) -> clause -> unit := "ltac2" "tac_eset". +Ltac2 @ external set : evar_flag -> ident option -> (unit -> constr) -> clause -> unit := "ltac2" "tac_set". Ltac2 @ external red : clause -> unit := "ltac2" "tac_red". Ltac2 @ external hnf : clause -> unit := "ltac2" "tac_hnf". @@ -108,18 +106,13 @@ Ltac2 @ external etransitivity : unit -> unit := "ltac2" "tac_etransitivity". Ltac2 @ external cut : constr -> unit := "ltac2" "tac_cut". -Ltac2 @ external left : bindings -> unit := "ltac2" "tac_left". -Ltac2 @ external eleft : bindings -> unit := "ltac2" "tac_eleft". -Ltac2 @ external right : bindings -> unit := "ltac2" "tac_right". -Ltac2 @ external eright : bindings -> unit := "ltac2" "tac_eright". +Ltac2 @ external left : evar_flag -> bindings -> unit := "ltac2" "tac_left". +Ltac2 @ external right : evar_flag -> bindings -> unit := "ltac2" "tac_right". -Ltac2 @ external constructor : unit -> unit := "ltac2" "tac_constructor". -Ltac2 @ external econstructor : unit -> unit := "ltac2" "tac_econstructor". -Ltac2 @ external split : bindings -> unit := "ltac2" "tac_split". -Ltac2 @ external esplit : bindings -> unit := "ltac2" "tac_esplit". +Ltac2 @ external constructor : evar_flag -> unit := "ltac2" "tac_constructor". +Ltac2 @ external split : evar_flag -> bindings -> unit := "ltac2" "tac_split". -Ltac2 @ external constructor_n : int -> bindings -> unit := "ltac2" "tac_constructorn". -Ltac2 @ external econstructor_n : int -> bindings -> unit := "ltac2" "tac_econstructorn". +Ltac2 @ external constructor_n : evar_flag -> int -> bindings -> unit := "ltac2" "tac_constructorn". Ltac2 @ external intros_until : hypothesis -> unit := "ltac2" "tac_introsuntil". -- cgit v1.2.3 From dbbefa2ed1f858c1a6de77672e3e1733ef4c28bf Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 2 Aug 2017 17:01:17 +0200 Subject: Code factorization in elim notation. --- theories/Notations.v | 20 ++++++-------------- 1 file changed, 6 insertions(+), 14 deletions(-) diff --git a/theories/Notations.v b/theories/Notations.v index 0487e324ca..2d7b4c8a8b 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -36,28 +36,20 @@ Ltac2 Notation "econstructor" := Std.constructor true. Ltac2 Notation "econstructor" n(tactic) bnd(bindings) := Std.constructor_n true n bnd. -Ltac2 eelim c bnd use := +Ltac2 elim0 ev c bnd use := let use := match use with | None => None | Some u => let ((_, c, wth)) := u in Some (c, wth) end in - Std.elim true (c, bnd) use. + Std.elim ev (c, bnd) use. -Ltac2 elim c bnd use := +Ltac2 Notation "elim" c(thunk(constr)) bnd(thunk(bindings)) + use(thunk(opt(seq("using", constr, bindings)))) := Control.with_holes (fun () => c (), bnd (), use ()) - (fun ((c, bnd, use)) => - let use := match use with - | None => None - | Some u => - let ((_, c, wth)) := u in Some (c, wth) - end in - Std.elim false (c, bnd) use). - -Ltac2 Notation "elim" c(thunk(constr)) bnd(thunk(bindings)) - use(thunk(opt(seq("using", constr, bindings)))) := elim c bnd use. + (fun ((c, bnd, use)) => elim0 false c bnd use). Ltac2 Notation "eelim" c(constr) bnd(bindings) use(opt(seq("using", constr, bindings))) := - eelim c bnd use. + elim0 true c bnd use. -- cgit v1.2.3 From d755c546a5c260232fd30971bd604b078d0afc18 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 2 Aug 2017 17:31:13 +0200 Subject: Properly implementing the notation to easily access hypotheses. --- doc/ltac2.md | 8 ++++++-- src/g_ltac2.ml4 | 10 ++++++++-- src/tac2quote.ml | 24 +++++++++++++++++++++++- src/tac2quote.mli | 6 ++++++ tests/example2.v | 7 +++++++ 5 files changed, 50 insertions(+), 5 deletions(-) diff --git a/doc/ltac2.md b/doc/ltac2.md index 12687e9aff..bf6d9eb583 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -470,8 +470,12 @@ as follows. constr:(fun x : nat => ltac2:(exact (hyp @x))) ``` -The `ltac2:(exact (hyp @x))` pattern is so common that we provide dedicated -Ltac2 and Coq term notations for it. +This pattern is so common that we provide dedicated Ltac2 and Coq term notations +for it. + +- `&x` as an Ltac2 expression expands to `hyp @x`. +- `&x` as an Coq constr expression expands to + `ltac2:(refine (fun () => hyp @x))`. #### Dynamic semantics diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index f558f9b9cc..bb98ea3e5d 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -136,6 +136,7 @@ GEXTEND Gram | id = Prim.qualid -> if Tac2env.is_constructor (snd id) then CTacCst (!@loc, RelId id) else CTacRef (RelId id) | "@"; id = Prim.ident -> Tac2quote.of_ident ~loc:!@loc id + | "&"; id = Prim.ident -> Tac2quote.of_hyp ~loc:!@loc id | "'"; c = Constr.constr -> inj_open_constr !@loc c | IDENT "constr"; ":"; "("; c = Constr.lconstr; ")" -> Tac2quote.of_constr ~loc:!@loc c | IDENT "open_constr"; ":"; "("; c = Constr.lconstr; ")" -> inj_open_constr !@loc c @@ -381,8 +382,13 @@ END GEXTEND Gram Pcoq.Constr.operconstr: LEVEL "0" [ [ IDENT "ltac2"; ":"; "("; tac = tac2expr; ")" -> - let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in - CAst.make ~loc:!@loc (CHole (None, IntroAnonymous, Some arg)) ] ] + let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in + CAst.make ~loc:!@loc (CHole (None, IntroAnonymous, Some arg)) + | "&"; id = Prim.ident -> + let tac = Tac2quote.of_exact_hyp ~loc:!@loc id in + let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in + CAst.make ~loc:!@loc (CHole (None, IntroAnonymous, Some arg)) + ] ] ; END diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 0e0a7b3fce..e30acc48ab 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -17,7 +17,13 @@ open Tac2core (** Syntactic quoting of expressions. *) -let std_core n = KerName.make2 Tac2env.std_prefix (Label.of_id (Id.of_string n)) +let control_prefix = + MPfile (DirPath.make (List.map Id.of_string ["Control"; "Ltac2"])) + +let kername prefix n = KerName.make2 prefix (Label.of_id (Id.of_string n)) +let std_core n = kername Tac2env.std_prefix n +let coq_core n = kername Tac2env.coq_prefix n +let control_core n = kername control_prefix n let dummy_loc = Loc.make_loc (-1, -1) @@ -114,3 +120,19 @@ and of_or_and_intro_pattern ?loc = function and of_intro_patterns ?loc l = of_list ?loc (List.map (of_intro_pattern ?loc) l) + +let of_hyp ?loc id = + let loc = Option.default dummy_loc loc in + let hyp = CTacRef (AbsKn (control_core "hyp")) in + CTacApp (loc, hyp, [of_ident ~loc id]) + +let thunk e = + let t_unit = coq_core "unit" in + let loc = Tac2intern.loc_of_tacexpr e in + let var = [CPatVar (Some loc, Anonymous), Some (CTypRef (loc, AbsKn (Other t_unit), []))] in + CTacFun (loc, var, e) + +let of_exact_hyp ?loc id = + let loc = Option.default dummy_loc loc in + let refine = CTacRef (AbsKn (control_core "refine")) in + CTacApp (loc, refine, [thunk (of_hyp ~loc id)]) diff --git a/src/tac2quote.mli b/src/tac2quote.mli index a311430a66..4cbe854f75 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -39,3 +39,9 @@ val of_bindings : ?loc:Loc.t -> bindings -> raw_tacexpr val of_intro_pattern : ?loc:Loc.t -> intro_pattern -> raw_tacexpr val of_intro_patterns : ?loc:Loc.t -> intro_pattern list -> raw_tacexpr + +val of_hyp : ?loc:Loc.t -> Id.t -> raw_tacexpr +(** id ↦ 'Control.hyp @id' *) + +val of_exact_hyp : ?loc:Loc.t -> Id.t -> raw_tacexpr +(** id ↦ 'Control.refine (fun () => Control.hyp @id') *) diff --git a/tests/example2.v b/tests/example2.v index 398f33561e..ca9e3dcff5 100644 --- a/tests/example2.v +++ b/tests/example2.v @@ -19,3 +19,10 @@ Proof. let myvar := Std.NamedHyp @x in split with (?myvar := 0). split. Qed. + +Goal (forall n : nat, n = 0 -> False) -> True. +Proof. +intros H. +elim &H with 0. +split. +Qed. -- cgit v1.2.3 From 7aab63c16dc5876f314208595b4b5d9d982ec1b1 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 2 Aug 2017 17:53:52 +0200 Subject: Fix compilation of horrible Ltac2 example. --- tests/stuff/ltac2.v | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/tests/stuff/ltac2.v b/tests/stuff/ltac2.v index 95bff7e569..370bc70d15 100644 --- a/tests/stuff/ltac2.v +++ b/tests/stuff/ltac2.v @@ -129,7 +129,7 @@ Abort. Goal (forall n : nat, n = 0 -> False) -> True. Proof. refine (fun () => '(fun H => _)). -Std.ecase (hyp @H, Std.ExplicitBindings [Std.NamedHyp @n, '0]). +Std.case true (hyp @H, Std.ExplicitBindings [Std.NamedHyp @n, '0]). refine (fun () => 'eq_refl). Qed. @@ -138,5 +138,6 @@ Proof. refine (fun () => '(fun x => _)). Std.cbv { Std.rBeta := true; Std.rMatch := true; Std.rFix := true; Std.rCofix := true; - Std.rZeta := true; Std.rDelta := false; Std.rConst := []; + Std.rZeta := true; Std.rDelta := true; Std.rConst := []; } { Std.on_hyps := None; Std.on_concl := Std.AllOccurrences }. +Abort. -- cgit v1.2.3 From 899476fa3dd2ae22f433a70fb860df0510a7ac88 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 2 Aug 2017 18:02:01 +0200 Subject: Expanding documentation. --- doc/ltac2.md | 78 ++++++++++++++++++++++++++++++++++++++++++++++++++---------- 1 file changed, 65 insertions(+), 13 deletions(-) diff --git a/doc/ltac2.md b/doc/ltac2.md index bf6d9eb583..20b043ea0b 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -44,7 +44,34 @@ We describe more in details each point in the remainder of this document. # ML component -The call-by-value functional language fragment is easy to implement. +## Overview + +Ltac2 is a member of the ML family of languages, in the sense that it is an +effectful call-by-value functional language, with static typing à la +Hindley-Milner. It is commonly accepted that ML constitutes a sweet spot in PL +design, as it is relatively expressive while not being either too lax +(contrarily to dynamic typing) nor too strict (contrarily to say, dependent +types). + +The main goal of Ltac2 is to serve as a meta-language for Coq. As such, it +naturally fits in the ML lineage, just as the historical ML was designed as +the tactic language for the LCF prover. It can also be seen as a general-purpose +language, by simply forgetting about the Coq-specific features. + +Sticking to a standard ML type system can be considered somewhat weak for a +meta-language designed to manipulate Coq terms. In particular, there is no +way to statically guarantee that a Coq term resulting from an Ltac2 +computation will be well-typed. This is actually a design choice, motivated +by retro-compatibility with Ltac1. Instead, well-typedness is deferred to +dynamic checks, allowing many primitive functions to fail whenever they are +provided with an ill-typed term. + +The language is naturally effectful as it manipulates the global state of the +proof engine. This allows to think of proof-modifying primitives as effects +in a straightforward way. Semantically, proof manipulation lives in a monad, +which allows to ensure that Ltac2 satisfies the same equations as a generic ML +with unspecified effects would do, e.g. function reduction is substitution +by a value. ## Type Syntax @@ -134,7 +161,8 @@ VERNAC ::= ## Term Syntax The syntax of the functional fragment is very close to the one of Ltac1, except -that it adds a true pattern-matching feature. +that it adds a true pattern-matching feature, as well as a few standard +constructions from ML. ``` VAR := LIDENT @@ -202,11 +230,21 @@ If the `RECFLAG` is set, the tactic is expanded into a recursive binding. ## Reduction We use the usual ML call-by-value reduction, with an otherwise unspecified -evaluation order. +evaluation order. This is a design choice making it compatible with OCaml, +if ever we implement native compilation. The expected equations are as follows. +``` +(fun x => t) V ≡ t{x := V} (βv) -Note that this is already a departure from Ltac1 which uses heuristic to -decide when evaluating an expression, e.g. the following do not evaluate the -same way. +let x := V in t ≡ t{x := V} (let) + +match C V₀ ... Vₙ with ... | C x₀ ... xₙ => t | ... end ≡ t {xᵢ := Vᵢ} (ι) + +(t any term, V values, C constructor) +``` + +Note that call-by-value reduction is already a departure from Ltac1 which uses +heuristics to decide when evaluating an expression. For instance, the following +expressions do not evaluate the same way in Ltac1. ``` foo (idtac; let x := 0 in bar) @@ -229,7 +267,7 @@ current hackish subtyping semantics, and one will have to resort to conversion functions. See notations though to make things more palatable. In this setting, all usual argument-free tactics have type `unit -> unit`, but -one can return as well a value of type `τ` thanks to terms of type `unit -> τ`, +one can return as well a value of type `t` thanks to terms of type `unit -> t`, or take additional arguments. ## Effects @@ -254,6 +292,8 @@ Intuitively a thunk of type `unit -> 'a` can do the following: - It can access a backtracking proof state, made out amongst other things of the current evar assignation and the list of goals under focus. +We describe more thoroughly the various effects existing in Ltac2 hereafter. + ### Standard IO The Ltac2 language features non-backtracking IO, notably mutable data and @@ -427,13 +467,16 @@ evaluation. - Evaluation is part of the dynamic semantics, i.e. it is done when a term gets effectively computed by Ltac2. +Remark that typing of Coq terms is a *dynamic* process occuring at Ltac2 +evaluation time, and not at Ltac2 typing time. + #### Static semantics During internalization, Coq variables are resolved and antiquotations are type-checked as Ltac2 terms, effectively producing a `glob_constr` in Coq implementation terminology. Note that although it went through the -type-checking of *Ltac2*, the resulting term has not been fully computed and -is potentially ill-typed as a Coq term. +type-checking of **Ltac2**, the resulting term has not been fully computed and +is potentially ill-typed as a runtime **Coq** term. ``` Ltac2 Definition myconstr () := constr:(nat -> 0). @@ -451,7 +494,7 @@ let x := '0 in constr:(1 + ltac2:(exact x)) Beware that the typing environment of typing of antiquotations is **not** expanded by the Coq binders from the term. Namely, it means that the following -expression will **not** type-check. +Ltac2 expression will **not** type-check. ``` constr:(fun x : nat => ltac2:(exact x)) // Error: Unbound variable 'x' @@ -462,6 +505,8 @@ not make sense in general. ``` constr:(fun x : nat => ltac2:(clear @x; exact x)) ``` +Indeed, a hypothesis can suddenly disappear from the runtime context if some +other tactic pulls the rug from under you. Rather, the tactic writer has to resort to the **dynamic** goal environment, and must write instead explicitly that she is accessing a hypothesis, typically @@ -482,9 +527,9 @@ for it. During evaluation, a quoted term is fully evaluated to a kernel term, and is in particular type-checked in the current environment. -Evaluation of a quoted term is described below. -- The quoted term is evaluated by the pretyper. -- Antiquotations are evaluated in a context where there is exactly one goal +Evaluation of a quoted term goes as follows. +- The quoted term is first evaluated by the pretyper. +- Antiquotations are then evaluated in a context where there is exactly one goal under focus, with the hypotheses coming from the current environment extended with the bound variables of the term, and the resulting term is fed into the quoted term. @@ -608,3 +653,10 @@ let c := fun () => constr:(nat -> nat) with ids := [@x; y] in Bar.f c ids Beware that the order of evaluation of multiple let-bindings is not specified, so that you may have to resort to thunking to ensure that side-effects are performed at the right time. + +# TODO + +- Implement deep pattern-matching. +- Implement compatibility layer with Ltac1 +- Craft an expressive set of primitive functions +- Implement native compilation to OCaml -- cgit v1.2.3 From 6e150eb19a55b16bbd4ea03964ee48f2d69084ed Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 2 Aug 2017 18:44:03 +0200 Subject: Typo in documentation. --- doc/ltac2.md | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/doc/ltac2.md b/doc/ltac2.md index 20b043ea0b..c2d930c9b6 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -642,7 +642,7 @@ Ltac2 Notation "foo" c(thunk(constr)) ids(list0(ident)) := Bar.f c ids. ``` Then the following expression ``` -let y := @X in foo (nat -> nat) x ?y +let y := @X in foo (nat -> nat) x $y ``` will expand at parsing time to ``` -- cgit v1.2.3 From 3007909ca1f65132bd0850d2be57e781e55707bd Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 2 Aug 2017 18:51:19 +0200 Subject: Tentatively implementing apply. --- _CoqProject | 2 ++ src/g_ltac2.ml4 | 5 ++++- src/ltac2_plugin.mlpack | 1 + src/tac2core.ml | 8 ++++++++ src/tac2entries.ml | 1 + src/tac2entries.mli | 1 + src/tac2stdlib.ml | 12 ++++++++++++ src/tac2tactics.ml | 25 +++++++++++++++++++++++++ src/tac2tactics.mli | 18 ++++++++++++++++++ tests/example2.v | 15 +++++++++++++++ theories/Notations.v | 35 ++++++++++++++++++++++++++++++----- theories/Std.v | 4 ++++ 12 files changed, 121 insertions(+), 6 deletions(-) create mode 100644 src/tac2tactics.ml create mode 100644 src/tac2tactics.mli diff --git a/_CoqProject b/_CoqProject index 583639612b..f202e1aed2 100644 --- a/_CoqProject +++ b/_CoqProject @@ -20,6 +20,8 @@ src/tac2core.mli src/tac2qexpr.mli src/tac2quote.ml src/tac2quote.mli +src/tac2tactics.ml +src/tac2tactics.mli src/tac2stdlib.ml src/tac2stdlib.mli src/g_ltac2.ml4 diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index bb98ea3e5d..ca3631799b 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -291,7 +291,7 @@ open Tac2entries.Pltac let loc_of_ne_list l = Loc.merge_opt (fst (List.hd l)) (fst (List.last l)) GEXTEND Gram - GLOBAL: q_ident q_bindings q_intropatterns; + GLOBAL: q_ident q_bindings q_intropattern q_intropatterns; ident_or_anti: [ [ id = Prim.ident -> QExpr id | LEFTQMARK; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) @@ -375,6 +375,9 @@ GEXTEND Gram q_intropatterns: [ [ ipat = intropatterns -> Tac2quote.of_intro_patterns ~loc:!@loc ipat ] ] ; + q_intropattern: + [ [ ipat = simple_intropattern -> Tac2quote.of_intro_pattern ~loc:!@loc ipat ] ] + ; END (** Extension of constr syntax *) diff --git a/src/ltac2_plugin.mlpack b/src/ltac2_plugin.mlpack index 8d2d7dc0f4..4c4082ad65 100644 --- a/src/ltac2_plugin.mlpack +++ b/src/ltac2_plugin.mlpack @@ -6,5 +6,6 @@ Tac2entries Tac2ffi Tac2core Tac2quote +Tac2tactics Tac2stdlib G_ltac2 diff --git a/src/tac2core.ml b/src/tac2core.ml index b45275210e..329c115be3 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -883,6 +883,14 @@ let () = add_scope "bindings" begin function | _ -> scope_fail () end +let () = add_scope "intropattern" begin function +| [] -> + let scope = Extend.Aentry Tac2entries.Pltac.q_intropattern in + let act tac = tac in + Tac2entries.ScopeRule (scope, act) +| _ -> scope_fail () +end + let () = add_scope "intropatterns" begin function | [] -> let scope = Extend.Aentry Tac2entries.Pltac.q_intropatterns in diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 0f32736096..52a5899d25 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -26,6 +26,7 @@ let tac2expr = Pcoq.Gram.entry_create "tactic:tac2expr" let q_ident = Pcoq.Gram.entry_create "tactic:q_ident" let q_bindings = Pcoq.Gram.entry_create "tactic:q_bindings" +let q_intropattern = Pcoq.Gram.entry_create "tactic:q_intropattern" let q_intropatterns = Pcoq.Gram.entry_create "tactic:q_intropatterns" end diff --git a/src/tac2entries.mli b/src/tac2entries.mli index e5031fdba2..2e51a4fb2e 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -59,5 +59,6 @@ val tac2expr : raw_tacexpr Pcoq.Gram.entry val q_ident : raw_tacexpr Pcoq.Gram.entry val q_bindings : raw_tacexpr Pcoq.Gram.entry +val q_intropattern : raw_tacexpr Pcoq.Gram.entry val q_intropatterns : raw_tacexpr Pcoq.Gram.entry end diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index f63252ec22..b678b65b82 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -12,6 +12,7 @@ open Misctypes open Genredexpr open Tac2expr open Tac2core +open Tac2tactics open Proofview.Notations module Value = Tac2ffi @@ -178,6 +179,17 @@ let () = define_prim2 "tac_intros" begin fun ev ipat -> Tactics.intros_patterns ev ipat end +let () = define_prim4 "tac_apply" begin fun adv ev cb ipat -> + let adv = Value.to_bool adv in + let ev = Value.to_bool ev in + let map_cb c = thaw c >>= fun c -> return (to_constr_with_bindings c) in + let cb = Value.to_list map_cb cb in + let map p = Value.to_option (fun p -> Loc.tag (to_intro_pattern p)) p in + let map_ipat p = to_pair Value.to_ident map p in + let ipat = Value.to_option map_ipat ipat in + Tac2tactics.apply adv ev cb ipat +end + let () = define_prim3 "tac_elim" begin fun ev c copt -> let ev = Value.to_bool ev in let c = to_constr_with_bindings c in diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml new file mode 100644 index 0000000000..2590d7daed --- /dev/null +++ b/src/tac2tactics.ml @@ -0,0 +1,25 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Tactics.apply_with_delayed_bindings_gen adv ev cb + | Some (id, cl) -> Tactics.apply_delayed_in adv ev id cb cl diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli new file mode 100644 index 0000000000..86278f177e --- /dev/null +++ b/src/tac2tactics.mli @@ -0,0 +1,18 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* evars_flag -> + EConstr.constr with_bindings tactic list -> + (Id.t * intro_pattern option) option -> unit tactic diff --git a/tests/example2.v b/tests/example2.v index ca9e3dcff5..fd5a9044e9 100644 --- a/tests/example2.v +++ b/tests/example2.v @@ -26,3 +26,18 @@ intros H. elim &H with 0. split. Qed. + +Goal forall (P : nat -> Prop), (forall n m, n = m -> P n) -> P 0. +Proof. +intros P H. +Fail apply &H. +apply &H with (m := 0). +split. +Qed. + +Goal forall (P : nat -> Prop), (forall n m, n = m -> P n) -> P 0. +Proof. +intros P H. +eapply &H. +split. +Qed. diff --git a/theories/Notations.v b/theories/Notations.v index 2d7b4c8a8b..1bc48d587a 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -37,11 +37,11 @@ Ltac2 Notation "econstructor" n(tactic) bnd(bindings) := Std.constructor_n true n bnd. Ltac2 elim0 ev c bnd use := - let use := match use with - | None => None - | Some u => - let ((_, c, wth)) := u in Some (c, wth) - end in + let use := match use with + | None => None + | Some u => + let ((_, c, wth)) := u in Some (c, wth) + end in Std.elim ev (c, bnd) use. Ltac2 Notation "elim" c(thunk(constr)) bnd(thunk(bindings)) @@ -53,3 +53,28 @@ Ltac2 Notation "elim" c(thunk(constr)) bnd(thunk(bindings)) Ltac2 Notation "eelim" c(constr) bnd(bindings) use(opt(seq("using", constr, bindings))) := elim0 true c bnd use. + +Ltac2 apply0 adv ev cb cl := + let cl := match cl with + | None => None + | Some p => + let ((_, id, ipat)) := p in + let p := match ipat with + | None => None + | Some p => + let ((_, ipat)) := p in + Some ipat + end in + Some (id, p) + end in + Std.apply adv ev cb cl. + +Ltac2 Notation "eapply" + cb(list1(thunk(seq(constr, bindings)), ",")) + cl(opt(seq(keyword("in"), ident, opt(seq(keyword("as"), intropattern))))) := + apply0 true true cb cl. + +Ltac2 Notation "apply" + cb(list1(thunk(seq(constr, bindings)), ",")) + cl(opt(seq(keyword("in"), ident, opt(seq(keyword("as"), intropattern))))) := + apply0 true false cb cl. diff --git a/theories/Std.v b/theories/Std.v index 20504f1247..3d1e8f462d 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -74,11 +74,15 @@ with or_and_intro_pattern := [ ]. Ltac2 Type evar_flag := bool. +Ltac2 Type advanced_flag := bool. (** Standard, built-in tactics. See Ltac1 for documentation. *) Ltac2 @ external intros : evar_flag -> intro_pattern list -> unit := "ltac2" "tac_intros". +Ltac2 @ external apply : advanced_flag -> evar_flag -> + (unit -> constr_with_bindings) list -> (ident * (intro_pattern option)) option -> unit := "ltac2" "tac_apply". + Ltac2 @ external elim : evar_flag -> constr_with_bindings -> constr_with_bindings option -> unit := "ltac2" "tac_elim". Ltac2 @ external case : evar_flag -> constr_with_bindings -> unit := "ltac2" "tac_case". -- cgit v1.2.3 From 9db02b3bfe35c15c9df8615f0e47a2a6407e858b Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 2 Aug 2017 20:37:12 +0200 Subject: Inserting enter functions in Ltac1 bindings. --- tests/example2.v | 20 ++++++++++++++ theories/Notations.v | 74 ++++++++++++++++++++++++++++++++-------------------- 2 files changed, 66 insertions(+), 28 deletions(-) diff --git a/tests/example2.v b/tests/example2.v index fd5a9044e9..79f230ab57 100644 --- a/tests/example2.v +++ b/tests/example2.v @@ -20,6 +20,13 @@ let myvar := Std.NamedHyp @x in split with (?myvar := 0). split. Qed. +Goal (forall n : nat, n = 0 -> False) -> True. +Proof. +intros H. +eelim &H. +split. +Qed. + Goal (forall n : nat, n = 0 -> False) -> True. Proof. intros H. @@ -41,3 +48,16 @@ intros P H. eapply &H. split. Qed. + +Goal exists n, n = 0. +Proof. +Fail constructor 1. +constructor 1 with (x := 0). +split. +Qed. + +Goal exists n, n = 0. +Proof. +econstructor 1. +split. +Qed. diff --git a/theories/Notations.v b/theories/Notations.v index 1bc48d587a..e7792c1555 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -9,49 +9,67 @@ Require Import Ltac2.Init. Require Ltac2.Control Ltac2.Std. -Ltac2 Notation "intros" p(intropatterns) := Std.intros false p. +(** Enter and check evar resolution *) +Ltac2 enter_h ev f arg := +match ev with +| true => Control.enter (fun () => f ev (arg ())) +| false => + Control.enter (fun () => + Control.with_holes arg (fun x => f ev x)) +end. -Ltac2 Notation "eintros" p(intropatterns) := Std.intros true p. +Ltac2 intros0 ev p := + Control.enter (fun () => Std.intros false p). -Ltac2 Notation "split" bnd(thunk(bindings)) := - Control.with_holes bnd (fun bnd => Std.split false bnd). +Ltac2 Notation "intros" p(intropatterns) := intros0 false p. -Ltac2 Notation "esplit" bnd(bindings) := Std.split true bnd. +Ltac2 Notation "eintros" p(intropatterns) := intros0 true p. -Ltac2 Notation "left" bnd(thunk(bindings)) := - Control.with_holes bnd (fun bnd => Std.left false bnd). +Ltac2 split0 ev bnd := + enter_h ev Std.split bnd. -Ltac2 Notation "eleft" bnd(bindings) := Std.left true bnd. +Ltac2 Notation "split" bnd(thunk(bindings)) := split0 false bnd. -Ltac2 Notation "right" bnd(thunk(bindings)) := - Control.with_holes bnd (fun bnd => Std.right false bnd). +Ltac2 Notation "esplit" bnd(thunk(bindings)) := split0 true bnd. -Ltac2 Notation "eright" bnd(bindings) := Std.right true bnd. +Ltac2 left0 ev bnd := enter_h ev Std.left bnd. -Ltac2 Notation "constructor" := Std.constructor false. -Ltac2 Notation "constructor" n(tactic) bnd(thunk(bindings)) := - Control.with_holes bnd (fun bnd => Std.constructor_n false n bnd). +Ltac2 Notation "left" bnd(thunk(bindings)) := left0 false bnd. -Ltac2 Notation "econstructor" := Std.constructor true. -Ltac2 Notation "econstructor" n(tactic) bnd(bindings) := - Std.constructor_n true n bnd. +Ltac2 Notation "eleft" bnd(thunk(bindings)) := left0 true bnd. + +Ltac2 right0 ev bnd := enter_h ev Std.right bnd. + +Ltac2 Notation "right" bnd(thunk(bindings)) := right0 false bnd. + +Ltac2 Notation "eright" bnd(thunk(bindings)) := right0 true bnd. + +Ltac2 constructor0 ev n bnd := + enter_h ev (fun ev bnd => Std.constructor_n ev n bnd) bnd. + +Ltac2 Notation "constructor" := Control.enter (fun () => Std.constructor false). +Ltac2 Notation "constructor" n(tactic) bnd(thunk(bindings)) := constructor0 false n bnd. + +Ltac2 Notation "econstructor" := Control.enter (fun () => Std.constructor true). +Ltac2 Notation "econstructor" n(tactic) bnd(thunk(bindings)) := constructor0 true n bnd. Ltac2 elim0 ev c bnd use := - let use := match use with - | None => None - | Some u => - let ((_, c, wth)) := u in Some (c, wth) - end in - Std.elim ev (c, bnd) use. + let f ev ((c, bnd, use)) := + let use := match use with + | None => None + | Some u => + let ((_, c, wth)) := u in Some (c, wth) + end in + Std.elim ev (c, bnd) use + in + enter_h ev f (fun () => c (), bnd (), use ()). Ltac2 Notation "elim" c(thunk(constr)) bnd(thunk(bindings)) use(thunk(opt(seq("using", constr, bindings)))) := - Control.with_holes - (fun () => c (), bnd (), use ()) - (fun ((c, bnd, use)) => elim0 false c bnd use). + elim0 false c bnd use. -Ltac2 Notation "eelim" c(constr) bnd(bindings) - use(opt(seq("using", constr, bindings))) := +Ltac2 Notation "eelim" c(thunk(constr)) bnd(thunk(bindings)) + use(thunk(opt(seq("using", constr, bindings)))) := elim0 true c bnd use. Ltac2 apply0 adv ev cb cl := -- cgit v1.2.3 From b84b03bb6230fca69cd9191ba0424402a5cd2330 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 3 Aug 2017 20:23:08 +0200 Subject: Introducing notations for destruct and induction arguments. --- src/g_ltac2.ml4 | 105 ++++++++++++++++++++++++++++++++++++++++++++++++++-- src/tac2core.ml | 8 ++++ src/tac2entries.ml | 1 + src/tac2entries.mli | 1 + src/tac2qexpr.mli | 23 ++++++++++++ src/tac2quote.ml | 79 ++++++++++++++++++++++++++++++++++++--- src/tac2quote.mli | 2 + tests/example2.v | 2 +- theories/Std.v | 13 +++++++ 9 files changed, 223 insertions(+), 11 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index ca3631799b..8c7db71a47 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -30,7 +30,7 @@ let test_lpar_idnum_coloneq = (match stream_nth 2 strm with | KEYWORD ":=" -> () | _ -> err ()) - | LEFTQMARK -> + | KEYWORD "$" -> (match stream_nth 2 strm with | IDENT _ -> (match stream_nth 3 strm with @@ -40,6 +40,20 @@ let test_lpar_idnum_coloneq = | _ -> err ()) | _ -> err ()) +(* Hack to recognize "(x)" *) +let test_lpar_id_rpar = + Gram.Entry.of_parser "lpar_id_coloneq" + (fun strm -> + match stream_nth 0 strm with + | KEYWORD "(" -> + (match stream_nth 1 strm with + | IDENT _ -> + (match stream_nth 2 strm with + | KEYWORD ")" -> () + | _ -> err ()) + | _ -> err ()) + | _ -> err ()) + let tac2expr = Tac2entries.Pltac.tac2expr let tac2type = Gram.entry_create "tactic:tac2type" let tac2def_val = Gram.entry_create "tactic:tac2def_val" @@ -291,17 +305,17 @@ open Tac2entries.Pltac let loc_of_ne_list l = Loc.merge_opt (fst (List.hd l)) (fst (List.last l)) GEXTEND Gram - GLOBAL: q_ident q_bindings q_intropattern q_intropatterns; + GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause; ident_or_anti: [ [ id = Prim.ident -> QExpr id - | LEFTQMARK; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) + | "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] ; q_ident: [ [ id = ident_or_anti -> Tac2quote.of_anti ~loc:!@loc Tac2quote.of_ident id ] ] ; simple_binding: - [ [ "("; LEFTQMARK; id = Prim.ident; ":="; c = Constr.lconstr; ")" -> + [ [ "("; "$"; id = Prim.ident; ":="; c = Constr.lconstr; ")" -> Loc.tag ~loc:!@loc (QAnti (Loc.tag ~loc:!@loc id), Tac2quote.of_open_constr ~loc:!@loc c) | "("; n = Prim.natural; ":="; c = Constr.lconstr; ")" -> Loc.tag ~loc:!@loc (QExpr (AnonHyp n), Tac2quote.of_open_constr ~loc:!@loc c) @@ -378,6 +392,89 @@ GEXTEND Gram q_intropattern: [ [ ipat = simple_intropattern -> Tac2quote.of_intro_pattern ~loc:!@loc ipat ] ] ; + nat_or_anti: + [ [ n = Prim.natural -> QExpr n + | "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) + ] ] + ; + eqn_ipat: + [ [ IDENT "eqn"; ":"; pat = naming_intropattern -> Some pat + | -> None + ] ] + ; + with_bindings: + [ [ "with"; bl = bindings -> bl | -> QNoBindings ] ] + ; + constr_with_bindings: + [ [ c = Constr.constr; l = with_bindings -> (c, l) ] ] + ; + destruction_arg: + [ [ n = Prim.natural -> QElimOnAnonHyp n + | (c, bnd) = constr_with_bindings -> QElimOnConstr (c, bnd) + ] ] + ; + as_or_and_ipat: + [ [ "as"; ipat = or_and_intropattern -> Some ipat + | -> None + ] ] + ; + occs_nums: + [ [ nl = LIST1 nat_or_anti -> QOnlyOccurrences nl + | "-"; n = nat_or_anti; nl = LIST0 nat_or_anti -> + QAllOccurrencesBut (n::nl) + ] ] + ; + occs: + [ [ "at"; occs = occs_nums -> occs | -> QAllOccurrences ] ] + ; + hypident: + [ [ id = ident_or_anti -> + id,Locus.InHyp + | "("; IDENT "type"; IDENT "of"; id = ident_or_anti; ")" -> + id,Locus.InHypTypeOnly + | "("; IDENT "value"; IDENT "of"; id = ident_or_anti; ")" -> + id,Locus.InHypValueOnly + ] ] + ; + hypident_occ: + [ [ (id,l)=hypident; occs=occs -> ((occs,id),l) ] ] + ; + in_clause: + [ [ "*"; occs=occs -> + { q_onhyps = None; q_concl_occs = occs } + | "*"; "|-"; occs = concl_occ -> + { q_onhyps = None; q_concl_occs = occs } + | hl = LIST0 hypident_occ SEP ","; "|-"; occs = concl_occ -> + { q_onhyps = Some hl; q_concl_occs = occs } + | hl = LIST0 hypident_occ SEP "," -> + { q_onhyps = Some hl; q_concl_occs = QNoOccurrences } + ] ] + ; + opt_clause: + [ [ "in"; cl = in_clause -> Some cl + | "at"; occs = occs_nums -> Some { q_onhyps = Some []; q_concl_occs = occs } + | -> None + ] ] + ; + concl_occ: + [ [ "*"; occs = occs -> occs + | -> QNoOccurrences + ] ] + ; + induction_clause: + [ [ c = destruction_arg; pat = as_or_and_ipat; eq = eqn_ipat; + cl = opt_clause -> + { + indcl_arg = c; + indcl_eqn = eq; + indcl_as = pat; + indcl_in = cl; + } + ] ] + ; + q_induction_clause: + [ [ cl = induction_clause -> Tac2quote.of_induction_clause ~loc:!@loc cl ] ] + ; END (** Extension of constr syntax *) diff --git a/src/tac2core.ml b/src/tac2core.ml index 329c115be3..45fa52ff9b 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -899,6 +899,14 @@ let () = add_scope "intropatterns" begin function | _ -> scope_fail () end +let () = add_scope "induction_clause" begin function +| [] -> + let scope = Extend.Aentry Tac2entries.Pltac.q_induction_clause in + let act tac = tac in + Tac2entries.ScopeRule (scope, act) +| _ -> scope_fail () +end + let () = add_generic_scope "constr" Pcoq.Constr.constr Stdarg.wit_constr let () = add_generic_scope "open_constr" Pcoq.Constr.constr Stdarg.wit_open_constr diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 52a5899d25..ce86e8aa33 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -28,6 +28,7 @@ let q_ident = Pcoq.Gram.entry_create "tactic:q_ident" let q_bindings = Pcoq.Gram.entry_create "tactic:q_bindings" let q_intropattern = Pcoq.Gram.entry_create "tactic:q_intropattern" let q_intropatterns = Pcoq.Gram.entry_create "tactic:q_intropatterns" +let q_induction_clause = Pcoq.Gram.entry_create "tactic:q_induction_clause" end (** Tactic definition *) diff --git a/src/tac2entries.mli b/src/tac2entries.mli index 2e51a4fb2e..1567551246 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -61,4 +61,5 @@ val q_ident : raw_tacexpr Pcoq.Gram.entry val q_bindings : raw_tacexpr Pcoq.Gram.entry val q_intropattern : raw_tacexpr Pcoq.Gram.entry val q_intropatterns : raw_tacexpr Pcoq.Gram.entry +val q_induction_clause : raw_tacexpr Pcoq.Gram.entry end diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index b68efe73ac..5075f2d7d4 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -40,3 +40,26 @@ and intro_pattern_action = and or_and_intro_pattern = | QIntroOrPattern of intro_pattern list list | QIntroAndPattern of intro_pattern list + +type occurrences = +| QAllOccurrences +| QAllOccurrencesBut of int or_anti list +| QNoOccurrences +| QOnlyOccurrences of int or_anti list + +type hyp_location = (occurrences * Id.t or_anti) * Locus.hyp_location_flag + +type clause = + { q_onhyps : hyp_location list option; q_concl_occs : occurrences; } + +type destruction_arg = +| QElimOnConstr of Constrexpr.constr_expr * bindings +| QElimOnIdent of Id.t +| QElimOnAnonHyp of int + +type induction_clause = { + indcl_arg : destruction_arg; + indcl_eqn : intro_pattern_naming option; + indcl_as : or_and_intro_pattern option; + indcl_in : clause option; +} diff --git a/src/tac2quote.ml b/src/tac2quote.ml index e30acc48ab..9858f611fe 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -36,13 +36,31 @@ let constructor ?loc kn args = let std_constructor ?loc name args = constructor ?loc (std_core name) args +let std_proj ?loc name = + AbsKn (std_core name) + +let thunk e = + let t_unit = coq_core "unit" in + let loc = Tac2intern.loc_of_tacexpr e in + let var = [CPatVar (Some loc, Anonymous), Some (CTypRef (loc, AbsKn (Other t_unit), []))] in + CTacFun (loc, var, e) + let of_pair ?loc (e1, e2) = let loc = Option.default dummy_loc loc in CTacApp (loc, CTacCst (loc, AbsKn (Tuple 2)), [e1; e2]) +let of_tuple ?loc el = + let loc = Option.default dummy_loc loc in + let len = List.length el in + CTacApp (loc, CTacCst (loc, AbsKn (Tuple len)), el) + let of_int ?loc n = CTacAtm (Loc.tag ?loc (AtmInt n)) +let of_option ?loc opt = match opt with +| None -> constructor ?loc (coq_core "None") [] +| Some e -> constructor ?loc (coq_core "Some") [e] + let inj_wit ?loc wit x = let loc = Option.default dummy_loc loc in CTacExt (loc, Genarg.in_gen (Genarg.rawwit wit) x) @@ -121,17 +139,66 @@ and of_or_and_intro_pattern ?loc = function and of_intro_patterns ?loc l = of_list ?loc (List.map (of_intro_pattern ?loc) l) +let of_hyp_location_flag ?loc = function +| Locus.InHyp -> std_constructor ?loc "InHyp" [] +| Locus.InHypTypeOnly -> std_constructor ?loc "InHypTypeOnly" [] +| Locus.InHypValueOnly -> std_constructor ?loc "InHypValueOnly" [] + +let of_occurrences ?loc occ = match occ with +| QAllOccurrences -> std_constructor ?loc "AllOccurrences" [] +| QAllOccurrencesBut occs -> + let map occ = of_anti ?loc of_int occ in + let occs = of_list ?loc (List.map map occs) in + std_constructor ?loc "AllOccurrencesBut" [occs] +| QNoOccurrences -> std_constructor ?loc "NoOccurrences" [] +| QOnlyOccurrences occs -> + let map occ = of_anti ?loc of_int occ in + let occs = of_list ?loc (List.map map occs) in + std_constructor ?loc "OnlyOccurrences" [occs] + +let of_hyp_location ?loc ((occs, id), flag) = + of_tuple ?loc [ + of_anti ?loc of_ident id; + of_occurrences ?loc occs; + of_hyp_location_flag ?loc flag; + ] + +let of_clause ?loc cl = + let loc = Option.default dummy_loc loc in + let hyps = of_option ~loc (Option.map (fun l -> of_list ~loc (List.map of_hyp_location l)) cl.q_onhyps) in + let concl = of_occurrences ~loc cl.q_concl_occs in + CTacRec (loc, [ + std_proj "on_hyps", hyps; + std_proj "on_concl", concl; + ]) + +let of_destruction_arg ?loc = function +| QElimOnConstr (c, bnd) -> + let c = of_constr ?loc c in + let bnd = of_bindings ?loc bnd in + let arg = thunk (of_pair ?loc (c, bnd)) in + std_constructor ?loc "ElimOnConstr" [arg] +| QElimOnIdent id -> std_constructor ?loc "ElimOnIdent" [of_ident ?loc id] +| QElimOnAnonHyp n -> std_constructor ?loc "ElimOnAnonHyp" [of_int ?loc n] + +let of_induction_clause ?loc cl = + let arg = of_destruction_arg ?loc cl.indcl_arg in + let eqn = of_option ?loc (Option.map of_intro_pattern_naming cl.indcl_eqn) in + let as_ = of_option ?loc (Option.map of_or_and_intro_pattern cl.indcl_as) in + let in_ = of_option ?loc (Option.map of_clause cl.indcl_in) in + let loc = Option.default dummy_loc loc in + CTacRec (loc, [ + std_proj "indcl_arg", arg; + std_proj "indcl_eqn", eqn; + std_proj "indcl_as", as_; + std_proj "indcl_in", in_; + ]) + let of_hyp ?loc id = let loc = Option.default dummy_loc loc in let hyp = CTacRef (AbsKn (control_core "hyp")) in CTacApp (loc, hyp, [of_ident ~loc id]) -let thunk e = - let t_unit = coq_core "unit" in - let loc = Tac2intern.loc_of_tacexpr e in - let var = [CPatVar (Some loc, Anonymous), Some (CTypRef (loc, AbsKn (Other t_unit), []))] in - CTacFun (loc, var, e) - let of_exact_hyp ?loc id = let loc = Option.default dummy_loc loc in let refine = CTacRef (AbsKn (control_core "refine")) in diff --git a/src/tac2quote.mli b/src/tac2quote.mli index 4cbe854f75..40ea58e334 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -40,6 +40,8 @@ val of_intro_pattern : ?loc:Loc.t -> intro_pattern -> raw_tacexpr val of_intro_patterns : ?loc:Loc.t -> intro_pattern list -> raw_tacexpr +val of_induction_clause : ?loc:Loc.t -> induction_clause -> raw_tacexpr + val of_hyp : ?loc:Loc.t -> Id.t -> raw_tacexpr (** id ↦ 'Control.hyp @id' *) diff --git a/tests/example2.v b/tests/example2.v index 79f230ab57..d89dcfd450 100644 --- a/tests/example2.v +++ b/tests/example2.v @@ -16,7 +16,7 @@ Qed. Goal exists n, n = 0. Proof. -let myvar := Std.NamedHyp @x in split with (?myvar := 0). +let myvar := Std.NamedHyp @x in split with ($myvar := 0). split. Qed. diff --git a/theories/Std.v b/theories/Std.v index 3d1e8f462d..c2027e41c7 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -73,6 +73,19 @@ with or_and_intro_pattern := [ | IntroAndPattern (intro_pattern list) ]. +Ltac2 Type destruction_arg := [ +| ElimOnConstr (unit -> constr_with_bindings) +| ElimOnIdent (ident) +| ElimOnAnonHyp (int) +]. + +Ltac2 Type induction_clause := { + indcl_arg : destruction_arg; + indcl_eqn : intro_pattern_naming option; + indcl_as : or_and_intro_pattern option; + indcl_in : clause option; +}. + Ltac2 Type evar_flag := bool. Ltac2 Type advanced_flag := bool. -- cgit v1.2.3 From fce4a1a9cbb57a636155181898ae4ecece5af59d Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 4 Aug 2017 13:04:10 +0200 Subject: Adding the induction and destruct tactics. --- src/tac2stdlib.ml | 41 ++++++++++++++++++++++++++++++++++++----- src/tac2tactics.ml | 21 +++++++++++++++++++++ src/tac2tactics.mli | 9 +++++++++ tests/example2.v | 34 ++++++++++++++++++++++++++++++++++ theories/Notations.v | 42 ++++++++++++++++++++++++++++++++++++++++++ theories/Std.v | 6 ++++++ 6 files changed, 148 insertions(+), 5 deletions(-) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index b678b65b82..7e421c8577 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -17,6 +17,10 @@ open Proofview.Notations module Value = Tac2ffi +let return x = Proofview.tclUNIT x +let v_unit = Value.of_unit () +let thaw f = Tac2interp.interp_app f [v_unit] + let to_pair f g = function | ValBlk (0, [| x; y |]) -> (f x, g y) | _ -> assert false @@ -119,13 +123,28 @@ and to_intro_patterns il = let map ipat = Loc.tag (to_intro_pattern ipat) in Value.to_list map il +let to_destruction_arg = function +| ValBlk (0, [| c |]) -> + let c = thaw c >>= fun c -> return (to_constr_with_bindings c) in + ElimOnConstr c +| ValBlk (1, [| id |]) -> ElimOnIdent (Loc.tag (Value.to_ident id)) +| ValBlk (2, [| n |]) -> ElimOnAnonHyp (Value.to_int n) +| _ -> assert false + +let to_induction_clause = function +| ValBlk (0, [| arg; eqn; as_; in_ |]) -> + let arg = to_destruction_arg arg in + let eqn = Value.to_option (fun p -> Loc.tag (to_intro_pattern_naming p)) eqn in + let as_ = Value.to_option (fun p -> Loc.tag (to_or_and_intro_pattern p)) as_ in + let in_ = Value.to_option to_clause in_ in + ((None, arg), eqn, as_, in_) +| _ -> + assert false + (** Standard tactics sharing their implementation with Ltac1 *) let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } -let return x = Proofview.tclUNIT x -let v_unit = Value.of_unit () - let lift tac = tac <*> return v_unit let wrap f = @@ -134,8 +153,6 @@ let wrap f = let wrap_unit f = return () >>= fun () -> f (); return v_unit -let thaw f = Tac2interp.interp_app f [v_unit] - let define_prim0 name tac = let tac = function | [_] -> lift tac @@ -245,6 +262,20 @@ let () = define_prim4 "tac_set" begin fun ev idopt c cl -> Tactics.letin_pat_tac ev None na (sigma, c) cl end +let () = define_prim3 "tac_destruct" begin fun ev ic using -> + let ev = Value.to_bool ev in + let ic = Value.to_list to_induction_clause ic in + let using = Value.to_option to_constr_with_bindings using in + Tac2tactics.induction_destruct false ev ic using +end + +let () = define_prim3 "tac_induction" begin fun ev ic using -> + let ev = Value.to_bool ev in + let ic = Value.to_list to_induction_clause ic in + let using = Value.to_option to_constr_with_bindings using in + Tac2tactics.induction_destruct true ev ic using +end + let () = define_prim1 "tac_red" begin fun cl -> let cl = to_clause cl in Tactics.reduce (Red false) cl diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index 2590d7daed..439250db78 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -6,7 +6,10 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open Misctypes +open Tactypes open Tactics +open Proofview open Tacmach.New open Tacticals.New open Proofview.Notations @@ -23,3 +26,21 @@ let apply adv ev cb cl = match cl with | None -> Tactics.apply_with_delayed_bindings_gen adv ev cb | Some (id, cl) -> Tactics.apply_delayed_in adv ev id cb cl + +type induction_clause = + EConstr.constr with_bindings tactic destruction_arg * + intro_pattern_naming option * + or_and_intro_pattern option * + Locus.clause option + +let map_destruction_arg = function +| ElimOnConstr c -> ElimOnConstr (delayed_of_tactic c) +| ElimOnIdent id -> ElimOnIdent id +| ElimOnAnonHyp n -> ElimOnAnonHyp n + +let map_induction_clause ((clear, arg), eqn, as_, occ) = + ((clear, map_destruction_arg arg), (eqn, as_), occ) + +let induction_destruct isrec ev ic using = + let ic = List.map map_induction_clause ic in + Tactics.induction_destruct isrec ev (ic, using) diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli index 86278f177e..f29793411a 100644 --- a/src/tac2tactics.mli +++ b/src/tac2tactics.mli @@ -16,3 +16,12 @@ open Proofview val apply : advanced_flag -> evars_flag -> EConstr.constr with_bindings tactic list -> (Id.t * intro_pattern option) option -> unit tactic + +type induction_clause = + EConstr.constr with_bindings tactic destruction_arg * + intro_pattern_naming option * + or_and_intro_pattern option * + Locus.clause option + +val induction_destruct : rec_flag -> evars_flag -> + induction_clause list -> EConstr.constr with_bindings option -> unit Proofview.tactic diff --git a/tests/example2.v b/tests/example2.v index d89dcfd450..812f9172c9 100644 --- a/tests/example2.v +++ b/tests/example2.v @@ -61,3 +61,37 @@ Proof. econstructor 1. split. Qed. + +Goal forall n, 0 + n = n. +Proof. +intros n. +induction &n as [|n] using nat_rect; split. +Qed. + +Goal forall n, 0 + n = n. +Proof. +intros n. +let n := @X in +let q := Std.NamedHyp @P in +induction &n as [|$n] using nat_rect with ($q := fun m => 0 + m = m); split. +Qed. + +Goal forall n, 0 + n = n. +Proof. +intros n. +destruct &n as [|n] using nat_rect; split. +Qed. + +Goal forall n, 0 + n = n. +Proof. +intros n. +let n := @X in +let q := Std.NamedHyp @P in +destruct &n as [|$n] using nat_rect with ($q := fun m => 0 + m = m); split. +Qed. + +Goal forall b1 b2, andb b1 b2 = andb b2 b1. +Proof. +intros b1 b2. +destruct &b1 as [|], &b2 as [|]; split. +Qed. diff --git a/theories/Notations.v b/theories/Notations.v index e7792c1555..20f01c3b48 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -96,3 +96,45 @@ Ltac2 Notation "apply" cb(list1(thunk(seq(constr, bindings)), ",")) cl(opt(seq(keyword("in"), ident, opt(seq(keyword("as"), intropattern))))) := apply0 true false cb cl. + +Ltac2 induction0 ev ic use := + let f ev use := + let use := match use with + | None => None + | Some u => + let ((_, c, wth)) := u in Some (c, wth) + end in + Std.induction ev ic use + in + enter_h ev f use. + +Ltac2 Notation "induction" + ic(list1(induction_clause, ",")) + use(thunk(opt(seq("using", constr, bindings)))) := + induction0 false ic use. + +Ltac2 Notation "einduction" + ic(list1(induction_clause, ",")) + use(thunk(opt(seq("using", constr, bindings)))) := + induction0 true ic use. + +Ltac2 destruct0 ev ic use := + let f ev use := + let use := match use with + | None => None + | Some u => + let ((_, c, wth)) := u in Some (c, wth) + end in + Std.destruct ev ic use + in + enter_h ev f use. + +Ltac2 Notation "destruct" + ic(list1(induction_clause, ",")) + use(thunk(opt(seq("using", constr, bindings)))) := + destruct0 false ic use. + +Ltac2 Notation "edestruct" + ic(list1(induction_clause, ",")) + use(thunk(opt(seq("using", constr, bindings)))) := + destruct0 true ic use. diff --git a/theories/Std.v b/theories/Std.v index c2027e41c7..19bdc4c82a 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -107,6 +107,12 @@ Ltac2 @ external enough : constr -> (unit -> unit) option option -> intro_patter Ltac2 @ external pose : ident option -> constr -> unit := "ltac2" "tac_pose". Ltac2 @ external set : evar_flag -> ident option -> (unit -> constr) -> clause -> unit := "ltac2" "tac_set". +Ltac2 @ external destruct : evar_flag -> induction_clause list -> + constr_with_bindings option -> unit := "ltac2" "tac_induction". + +Ltac2 @ external induction : evar_flag -> induction_clause list -> + constr_with_bindings option -> unit := "ltac2" "tac_induction". + Ltac2 @ external red : clause -> unit := "ltac2" "tac_red". Ltac2 @ external hnf : clause -> unit := "ltac2" "tac_hnf". Ltac2 @ external cbv : red_flags -> clause -> unit := "ltac2" "tac_cbv". -- cgit v1.2.3 From 2e01ea9e1ab0f9e8d90dd4e4ac598bc1691b9272 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 4 Aug 2017 15:48:07 +0200 Subject: More precise type for quoted structures. --- src/g_ltac2.ml4 | 7 +++---- src/tac2qexpr.mli | 4 ++-- src/tac2quote.ml | 6 ++++-- 3 files changed, 9 insertions(+), 8 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 8c7db71a47..bfef4fab8d 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -316,18 +316,17 @@ GEXTEND Gram ; simple_binding: [ [ "("; "$"; id = Prim.ident; ":="; c = Constr.lconstr; ")" -> - Loc.tag ~loc:!@loc (QAnti (Loc.tag ~loc:!@loc id), Tac2quote.of_open_constr ~loc:!@loc c) + Loc.tag ~loc:!@loc (QAnti (Loc.tag ~loc:!@loc id), c) | "("; n = Prim.natural; ":="; c = Constr.lconstr; ")" -> - Loc.tag ~loc:!@loc (QExpr (AnonHyp n), Tac2quote.of_open_constr ~loc:!@loc c) + Loc.tag ~loc:!@loc (QExpr (AnonHyp n), c) | "("; id = Prim.ident; ":="; c = Constr.lconstr; ")" -> - Loc.tag ~loc:!@loc (QExpr (NamedHyp id), Tac2quote.of_open_constr ~loc:!@loc c) + Loc.tag ~loc:!@loc (QExpr (NamedHyp id), c) ] ] ; bindings: [ [ test_lpar_idnum_coloneq; bl = LIST1 simple_binding -> QExplicitBindings bl | bl = LIST1 Constr.constr -> - let bl = List.map (fun c -> Tac2quote.of_open_constr ~loc:!@loc c) bl in QImplicitBindings bl ] ] ; diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index 5075f2d7d4..8a61590a1d 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -19,8 +19,8 @@ type 'a or_anti = | QAnti of Id.t located type bindings = -| QImplicitBindings of raw_tacexpr list -| QExplicitBindings of (Misctypes.quantified_hypothesis or_anti * raw_tacexpr) Loc.located list +| QImplicitBindings of Constrexpr.constr_expr list +| QExplicitBindings of (Misctypes.quantified_hypothesis or_anti * Constrexpr.constr_expr) Loc.located list | QNoBindings type intro_pattern = diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 9858f611fe..a053bd799f 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -98,9 +98,11 @@ let of_bindings ?loc = function | QNoBindings -> std_constructor ?loc "NoBindings" [] | QImplicitBindings tl -> + let tl = List.map (fun c -> of_open_constr ?loc c) tl in std_constructor ?loc "ImplicitBindings" [of_list ?loc tl] | QExplicitBindings tl -> - let tl = List.map (fun (loc, (qhyp, e)) -> of_pair ?loc (of_anti ?loc of_qhyp qhyp, e)) tl in + let map (loc, (qhyp, e)) = of_pair ?loc (of_anti ?loc of_qhyp qhyp, of_open_constr ?loc e) in + let tl = List.map map tl in std_constructor ?loc "ExplicitBindings" [of_list ?loc tl] let rec of_intro_pattern ?loc = function @@ -174,7 +176,7 @@ let of_clause ?loc cl = let of_destruction_arg ?loc = function | QElimOnConstr (c, bnd) -> - let c = of_constr ?loc c in + let c = of_open_constr ?loc c in let bnd = of_bindings ?loc bnd in let arg = thunk (of_pair ?loc (c, bnd)) in std_constructor ?loc "ElimOnConstr" [arg] -- cgit v1.2.3 From 8bf0f3383fcde637ed9363f080d875a9ef0a138f Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 4 Aug 2017 15:52:37 +0200 Subject: Adding locations to quotation types. --- src/g_ltac2.ml4 | 125 +++++++++++++++++++++++++++++++++--------------------- src/tac2qexpr.mli | 49 +++++++++++++-------- src/tac2quote.ml | 66 ++++++++++++++-------------- src/tac2quote.mli | 25 +++++------ 4 files changed, 156 insertions(+), 109 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index bfef4fab8d..b10f8d66bd 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -149,12 +149,12 @@ GEXTEND Gram | s = Prim.string -> CTacAtm (Loc.tag ~loc:!@loc (AtmStr s)) | id = Prim.qualid -> if Tac2env.is_constructor (snd id) then CTacCst (!@loc, RelId id) else CTacRef (RelId id) - | "@"; id = Prim.ident -> Tac2quote.of_ident ~loc:!@loc id - | "&"; id = Prim.ident -> Tac2quote.of_hyp ~loc:!@loc id + | "@"; id = Prim.ident -> Tac2quote.of_ident (Loc.tag ~loc:!@loc id) + | "&"; id = lident -> Tac2quote.of_hyp ~loc:!@loc id | "'"; c = Constr.constr -> inj_open_constr !@loc c - | IDENT "constr"; ":"; "("; c = Constr.lconstr; ")" -> Tac2quote.of_constr ~loc:!@loc c - | IDENT "open_constr"; ":"; "("; c = Constr.lconstr; ")" -> inj_open_constr !@loc c - | IDENT "ident"; ":"; "("; c = Prim.ident; ")" -> Tac2quote.of_ident ~loc:!@loc c + | IDENT "constr"; ":"; "("; c = Constr.lconstr; ")" -> Tac2quote.of_constr c + | IDENT "open_constr"; ":"; "("; c = Constr.lconstr; ")" -> Tac2quote.of_open_constr c + | IDENT "ident"; ":"; "("; c = lident; ")" -> Tac2quote.of_ident c | IDENT "pattern"; ":"; "("; c = Constr.lconstr_pattern; ")" -> inj_pattern !@loc c ] ] ; @@ -296,6 +296,9 @@ GEXTEND Gram StrSyn (toks, n, e) ] ] ; + lident: + [ [ id = Prim.ident -> Loc.tag ~loc:!@loc id ] ] + ; END (** Quotation scopes used by notations *) @@ -306,71 +309,92 @@ let loc_of_ne_list l = Loc.merge_opt (fst (List.hd l)) (fst (List.last l)) GEXTEND Gram GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause; + anti: + [ [ "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] + ; ident_or_anti: - [ [ id = Prim.ident -> QExpr id + [ [ id = lident -> QExpr id | "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] ; + lident: + [ [ id = Prim.ident -> Loc.tag ~loc:!@loc id ] ] + ; + lnatural: + [ [ n = Prim.natural -> Loc.tag ~loc:!@loc n ] ] + ; q_ident: [ [ id = ident_or_anti -> Tac2quote.of_anti ~loc:!@loc Tac2quote.of_ident id ] ] ; + qhyp: + [ [ x = anti -> x + | n = lnatural -> QExpr (Loc.tag ~loc:!@loc @@ QAnonHyp n) + | id = lident -> QExpr (Loc.tag ~loc:!@loc @@ QNamedHyp id) + ] ] + ; simple_binding: - [ [ "("; "$"; id = Prim.ident; ":="; c = Constr.lconstr; ")" -> - Loc.tag ~loc:!@loc (QAnti (Loc.tag ~loc:!@loc id), c) - | "("; n = Prim.natural; ":="; c = Constr.lconstr; ")" -> - Loc.tag ~loc:!@loc (QExpr (AnonHyp n), c) - | "("; id = Prim.ident; ":="; c = Constr.lconstr; ")" -> - Loc.tag ~loc:!@loc (QExpr (NamedHyp id), c) + [ [ "("; h = qhyp; ":="; c = Constr.lconstr; ")" -> + Loc.tag ~loc:!@loc (h, c) ] ] ; bindings: [ [ test_lpar_idnum_coloneq; bl = LIST1 simple_binding -> - QExplicitBindings bl + Loc.tag ~loc:!@loc @@ QExplicitBindings bl | bl = LIST1 Constr.constr -> - QImplicitBindings bl + Loc.tag ~loc:!@loc @@ QImplicitBindings bl ] ] ; q_bindings: - [ [ "with"; bl = bindings -> Tac2quote.of_bindings ~loc:!@loc bl - | -> Tac2quote.of_bindings ~loc:!@loc QNoBindings - ] ] + [ [ bl = with_bindings -> Tac2quote.of_bindings bl ] ] ; intropatterns: - [ [ l = LIST0 nonsimple_intropattern -> l ]] + [ [ l = LIST0 nonsimple_intropattern -> Loc.tag ~loc:!@loc l ]] ; (* ne_intropatterns: *) (* [ [ l = LIST1 nonsimple_intropattern -> l ]] *) (* ; *) or_and_intropattern: - [ [ "["; tc = LIST1 intropatterns SEP "|"; "]" -> QIntroOrPattern tc - | "()" -> QIntroAndPattern [] - | "("; si = simple_intropattern; ")" -> QIntroAndPattern [si] + [ [ "["; tc = LIST1 intropatterns SEP "|"; "]" -> Loc.tag ~loc:!@loc @@ QIntroOrPattern tc + | "()" -> Loc.tag ~loc:!@loc @@ QIntroAndPattern (Loc.tag ~loc:!@loc []) + | "("; si = simple_intropattern; ")" -> Loc.tag ~loc:!@loc @@ QIntroAndPattern (Loc.tag ~loc:!@loc [si]) | "("; si = simple_intropattern; ","; tc = LIST1 simple_intropattern SEP "," ; ")" -> - QIntroAndPattern (si::tc) + Loc.tag ~loc:!@loc @@ QIntroAndPattern (Loc.tag ~loc:!@loc (si::tc)) | "("; si = simple_intropattern; "&"; tc = LIST1 simple_intropattern SEP "&" ; ")" -> (* (A & B & C) is translated into (A,(B,C)) *) let rec pairify = function - | ([]|[_]|[_;_]) as l -> l - | t::q -> [t; (QIntroAction (QIntroOrAndPattern (QIntroAndPattern (pairify q))))] - in QIntroAndPattern (pairify (si::tc)) ] ] + | ([]|[_]|[_;_]) as l -> Loc.tag ~loc:!@loc l + | t::q -> + let q = + Loc.tag ~loc:!@loc @@ + QIntroAction (Loc.tag ~loc:!@loc @@ + QIntroOrAndPattern (Loc.tag ~loc:!@loc @@ + QIntroAndPattern (pairify q))) + in + Loc.tag ~loc:!@loc [t; q] + in Loc.tag ~loc:!@loc @@ QIntroAndPattern (pairify (si::tc)) ] ] ; equality_intropattern: - [ [ "->" -> QIntroRewrite true - | "<-" -> QIntroRewrite false - | "[="; tc = intropatterns; "]" -> QIntroInjection tc ] ] + [ [ "->" -> Loc.tag ~loc:!@loc @@ QIntroRewrite true + | "<-" -> Loc.tag ~loc:!@loc @@ QIntroRewrite false + | "[="; tc = intropatterns; "]" -> Loc.tag ~loc:!@loc @@ QIntroInjection tc ] ] ; naming_intropattern: - [ [ LEFTQMARK; id = Prim.ident -> QIntroFresh (QExpr id) - | "?$"; id = Prim.ident -> QIntroFresh (QAnti (Loc.tag ~loc:!@loc id)) - | "?" -> QIntroAnonymous - | id = ident_or_anti -> QIntroIdentifier id ] ] + [ [ LEFTQMARK; id = lident -> + Loc.tag ~loc:!@loc @@ QIntroFresh (QExpr id) + | "?$"; id = lident -> + Loc.tag ~loc:!@loc @@ QIntroFresh (QAnti id) + | "?" -> + Loc.tag ~loc:!@loc @@ QIntroAnonymous + | id = ident_or_anti -> + Loc.tag ~loc:!@loc @@ QIntroIdentifier id + ] ] ; nonsimple_intropattern: [ [ l = simple_intropattern -> l - | "*" -> QIntroForthcoming true - | "**" -> QIntroForthcoming false ]] + | "*" -> Loc.tag ~loc:!@loc @@ QIntroForthcoming true + | "**" -> Loc.tag ~loc:!@loc @@ QIntroForthcoming false ]] ; simple_intropattern: [ [ pat = simple_intropattern_closed -> @@ -380,19 +404,24 @@ GEXTEND Gram ] ] ; simple_intropattern_closed: - [ [ pat = or_and_intropattern -> QIntroAction (QIntroOrAndPattern pat) - | pat = equality_intropattern -> QIntroAction pat - | "_" -> QIntroAction QIntroWildcard - | pat = naming_intropattern -> QIntroNaming pat ] ] + [ [ pat = or_and_intropattern -> + Loc.tag ~loc:!@loc @@ QIntroAction (Loc.tag ~loc:!@loc @@ QIntroOrAndPattern pat) + | pat = equality_intropattern -> + Loc.tag ~loc:!@loc @@ QIntroAction pat + | "_" -> + Loc.tag ~loc:!@loc @@ QIntroAction (Loc.tag ~loc:!@loc @@ QIntroWildcard) + | pat = naming_intropattern -> + Loc.tag ~loc:!@loc @@ QIntroNaming pat + ] ] ; q_intropatterns: - [ [ ipat = intropatterns -> Tac2quote.of_intro_patterns ~loc:!@loc ipat ] ] + [ [ ipat = intropatterns -> Tac2quote.of_intro_patterns ipat ] ] ; q_intropattern: - [ [ ipat = simple_intropattern -> Tac2quote.of_intro_pattern ~loc:!@loc ipat ] ] + [ [ ipat = simple_intropattern -> Tac2quote.of_intro_pattern ipat ] ] ; nat_or_anti: - [ [ n = Prim.natural -> QExpr n + [ [ n = lnatural -> QExpr n | "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] ; @@ -402,14 +431,14 @@ GEXTEND Gram ] ] ; with_bindings: - [ [ "with"; bl = bindings -> bl | -> QNoBindings ] ] + [ [ "with"; bl = bindings -> bl | -> Loc.tag ~loc:!@loc @@ QNoBindings ] ] ; constr_with_bindings: - [ [ c = Constr.constr; l = with_bindings -> (c, l) ] ] + [ [ c = Constr.constr; l = with_bindings -> Loc.tag ~loc:!@loc @@ (c, l) ] ] ; destruction_arg: - [ [ n = Prim.natural -> QElimOnAnonHyp n - | (c, bnd) = constr_with_bindings -> QElimOnConstr (c, bnd) + [ [ n = lnatural -> QElimOnAnonHyp n + | c = constr_with_bindings -> QElimOnConstr c ] ] ; as_or_and_ipat: @@ -463,7 +492,7 @@ GEXTEND Gram induction_clause: [ [ c = destruction_arg; pat = as_or_and_ipat; eq = eqn_ipat; cl = opt_clause -> - { + Loc.tag ~loc:!@loc @@ { indcl_arg = c; indcl_eqn = eq; indcl_as = pat; @@ -472,7 +501,7 @@ GEXTEND Gram ] ] ; q_induction_clause: - [ [ cl = induction_clause -> Tac2quote.of_induction_clause ~loc:!@loc cl ] ] + [ [ cl = induction_clause -> Tac2quote.of_induction_clause cl ] ] ; END @@ -483,7 +512,7 @@ GEXTEND Gram [ [ IDENT "ltac2"; ":"; "("; tac = tac2expr; ")" -> let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in CAst.make ~loc:!@loc (CHole (None, IntroAnonymous, Some arg)) - | "&"; id = Prim.ident -> + | "&"; id = [ id = Prim.ident -> Loc.tag ~loc:!@loc id ] -> let tac = Tac2quote.of_exact_hyp ~loc:!@loc id in let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in CAst.make ~loc:!@loc (CHole (None, IntroAnonymous, Some arg)) diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index 8a61590a1d..0eb9e9f4b5 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -18,48 +18,61 @@ type 'a or_anti = | QExpr of 'a | QAnti of Id.t located -type bindings = +type quantified_hypothesis = +| QAnonHyp of int located +| QNamedHyp of Id.t located + +type bindings_r = | QImplicitBindings of Constrexpr.constr_expr list -| QExplicitBindings of (Misctypes.quantified_hypothesis or_anti * Constrexpr.constr_expr) Loc.located list +| QExplicitBindings of (quantified_hypothesis located or_anti * Constrexpr.constr_expr) located list | QNoBindings -type intro_pattern = +type bindings = bindings_r located + +type intro_pattern_r = | QIntroForthcoming of bool | QIntroNaming of intro_pattern_naming | QIntroAction of intro_pattern_action -and intro_pattern_naming = -| QIntroIdentifier of Id.t or_anti -| QIntroFresh of Id.t or_anti +and intro_pattern_naming_r = +| QIntroIdentifier of Id.t located or_anti +| QIntroFresh of Id.t located or_anti | QIntroAnonymous -and intro_pattern_action = +and intro_pattern_action_r = | QIntroWildcard | QIntroOrAndPattern of or_and_intro_pattern -| QIntroInjection of intro_pattern list +| QIntroInjection of intro_pattern list located (* | QIntroApplyOn of Empty.t (** Not implemented yet *) *) | QIntroRewrite of bool -and or_and_intro_pattern = -| QIntroOrPattern of intro_pattern list list -| QIntroAndPattern of intro_pattern list +and or_and_intro_pattern_r = +| QIntroOrPattern of intro_pattern list located list +| QIntroAndPattern of intro_pattern list located + +and intro_pattern = intro_pattern_r located +and intro_pattern_naming = intro_pattern_naming_r located +and intro_pattern_action = intro_pattern_action_r located +and or_and_intro_pattern = or_and_intro_pattern_r located type occurrences = | QAllOccurrences -| QAllOccurrencesBut of int or_anti list +| QAllOccurrencesBut of int located or_anti list | QNoOccurrences -| QOnlyOccurrences of int or_anti list +| QOnlyOccurrences of int located or_anti list -type hyp_location = (occurrences * Id.t or_anti) * Locus.hyp_location_flag +type hyp_location = (occurrences * Id.t located or_anti) * Locus.hyp_location_flag type clause = { q_onhyps : hyp_location list option; q_concl_occs : occurrences; } type destruction_arg = -| QElimOnConstr of Constrexpr.constr_expr * bindings -| QElimOnIdent of Id.t -| QElimOnAnonHyp of int +| QElimOnConstr of (Constrexpr.constr_expr * bindings) located +| QElimOnIdent of Id.t located +| QElimOnAnonHyp of int located -type induction_clause = { +type induction_clause_r = { indcl_arg : destruction_arg; indcl_eqn : intro_pattern_naming option; indcl_as : or_and_intro_pattern option; indcl_in : clause option; } + +type induction_clause = induction_clause_r located diff --git a/src/tac2quote.ml b/src/tac2quote.ml index a053bd799f..1cba488768 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -54,7 +54,7 @@ let of_tuple ?loc el = let len = List.length el in CTacApp (loc, CTacCst (loc, AbsKn (Tuple len)), el) -let of_int ?loc n = +let of_int (loc, n) = CTacAtm (Loc.tag ?loc (AtmInt n)) let of_option ?loc opt = match opt with @@ -65,21 +65,25 @@ let inj_wit ?loc wit x = let loc = Option.default dummy_loc loc in CTacExt (loc, Genarg.in_gen (Genarg.rawwit wit) x) -let of_variable ?loc id = +let of_variable (loc, id) = let qid = Libnames.qualid_of_ident id in if Tac2env.is_constructor qid then CErrors.user_err ?loc (str "Invalid identifier") else CTacRef (RelId (Loc.tag ?loc qid)) let of_anti ?loc f = function -| QExpr x -> f ?loc x -| QAnti (loc, id) -> of_variable ?loc id +| QExpr x -> f x +| QAnti id -> of_variable id -let of_ident ?loc id = inj_wit ?loc Stdarg.wit_ident id +let of_ident (loc, id) = inj_wit ?loc Stdarg.wit_ident id -let of_constr ?loc c = inj_wit ?loc Stdarg.wit_constr c +let of_constr c = + let loc = Constrexpr_ops.constr_loc c in + inj_wit ?loc Stdarg.wit_constr c -let of_open_constr ?loc c = inj_wit ?loc Stdarg.wit_open_constr c +let of_open_constr c = + let loc = Constrexpr_ops.constr_loc c in + inj_wit ?loc Stdarg.wit_open_constr c let of_bool ?loc b = let c = if b then Core.c_true else Core.c_false in @@ -90,22 +94,22 @@ let rec of_list ?loc = function | e :: l -> constructor ?loc Core.c_cons [e; of_list ?loc l] -let of_qhyp ?loc = function -| AnonHyp n -> constructor Core.c_anon_hyp [of_int ?loc n] -| NamedHyp id -> constructor Core.c_named_hyp [of_ident ?loc id] +let of_qhyp (loc, h) = match h with +| QAnonHyp n -> constructor ?loc Core.c_anon_hyp [of_int n] +| QNamedHyp id -> constructor ?loc Core.c_named_hyp [of_ident id] -let of_bindings ?loc = function +let of_bindings (loc, b) = match b with | QNoBindings -> std_constructor ?loc "NoBindings" [] | QImplicitBindings tl -> - let tl = List.map (fun c -> of_open_constr ?loc c) tl in + let tl = List.map of_open_constr tl in std_constructor ?loc "ImplicitBindings" [of_list ?loc tl] | QExplicitBindings tl -> - let map (loc, (qhyp, e)) = of_pair ?loc (of_anti ?loc of_qhyp qhyp, of_open_constr ?loc e) in + let map (loc, (qhyp, e)) = of_pair ?loc (of_anti ?loc of_qhyp qhyp, of_open_constr e) in let tl = List.map map tl in std_constructor ?loc "ExplicitBindings" [of_list ?loc tl] -let rec of_intro_pattern ?loc = function +let rec of_intro_pattern (loc, pat) = match pat with | QIntroForthcoming b -> std_constructor ?loc "IntroForthcoming" [of_bool b] | QIntroNaming iname -> @@ -113,33 +117,33 @@ let rec of_intro_pattern ?loc = function | QIntroAction iact -> std_constructor ?loc "IntroAction" [of_intro_pattern_action iact] -and of_intro_pattern_naming ?loc = function +and of_intro_pattern_naming (loc, pat) = match pat with | QIntroIdentifier id -> - std_constructor ?loc "IntroIdentifier" [of_anti ?loc of_ident id] + std_constructor ?loc "IntroIdentifier" [of_anti of_ident id] | QIntroFresh id -> std_constructor ?loc "IntroFresh" [of_anti ?loc of_ident id] | QIntroAnonymous -> std_constructor ?loc "IntroAnonymous" [] -and of_intro_pattern_action ?loc = function +and of_intro_pattern_action (loc, pat) = match pat with | QIntroWildcard -> std_constructor ?loc "IntroWildcard" [] | QIntroOrAndPattern pat -> - std_constructor ?loc "IntroOrAndPattern" [of_or_and_intro_pattern ?loc pat] + std_constructor ?loc "IntroOrAndPattern" [of_or_and_intro_pattern pat] | QIntroInjection il -> - std_constructor ?loc "IntroInjection" [of_intro_patterns ?loc il] + std_constructor ?loc "IntroInjection" [of_intro_patterns il] | QIntroRewrite b -> std_constructor ?loc "IntroRewrite" [of_bool ?loc b] -and of_or_and_intro_pattern ?loc = function +and of_or_and_intro_pattern (loc, pat) = match pat with | QIntroOrPattern ill -> - let ill = List.map (of_intro_patterns ?loc) ill in + let ill = List.map of_intro_patterns ill in std_constructor ?loc "IntroOrPattern" [of_list ?loc ill] | QIntroAndPattern il -> - std_constructor ?loc "IntroAndPattern" [of_intro_patterns ?loc il] + std_constructor ?loc "IntroAndPattern" [of_intro_patterns il] -and of_intro_patterns ?loc l = - of_list ?loc (List.map (of_intro_pattern ?loc) l) +and of_intro_patterns (loc, l) = + of_list ?loc (List.map of_intro_pattern l) let of_hyp_location_flag ?loc = function | Locus.InHyp -> std_constructor ?loc "InHyp" [] @@ -175,15 +179,15 @@ let of_clause ?loc cl = ]) let of_destruction_arg ?loc = function -| QElimOnConstr (c, bnd) -> - let c = of_open_constr ?loc c in - let bnd = of_bindings ?loc bnd in +| QElimOnConstr (loc, (c, bnd)) -> + let c = of_open_constr c in + let bnd = of_bindings bnd in let arg = thunk (of_pair ?loc (c, bnd)) in std_constructor ?loc "ElimOnConstr" [arg] -| QElimOnIdent id -> std_constructor ?loc "ElimOnIdent" [of_ident ?loc id] -| QElimOnAnonHyp n -> std_constructor ?loc "ElimOnAnonHyp" [of_int ?loc n] +| QElimOnIdent id -> std_constructor ?loc "ElimOnIdent" [of_ident id] +| QElimOnAnonHyp n -> std_constructor ?loc "ElimOnAnonHyp" [of_int n] -let of_induction_clause ?loc cl = +let of_induction_clause (loc, cl) = let arg = of_destruction_arg ?loc cl.indcl_arg in let eqn = of_option ?loc (Option.map of_intro_pattern_naming cl.indcl_eqn) in let as_ = of_option ?loc (Option.map of_or_and_intro_pattern cl.indcl_as) in @@ -199,7 +203,7 @@ let of_induction_clause ?loc cl = let of_hyp ?loc id = let loc = Option.default dummy_loc loc in let hyp = CTacRef (AbsKn (control_core "hyp")) in - CTacApp (loc, hyp, [of_ident ~loc id]) + CTacApp (loc, hyp, [of_ident id]) let of_exact_hyp ?loc id = let loc = Option.default dummy_loc loc in diff --git a/src/tac2quote.mli b/src/tac2quote.mli index 40ea58e334..404e9378e0 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open Loc open Names open Misctypes open Tac2qexpr @@ -18,32 +19,32 @@ open Tac2expr val constructor : ?loc:Loc.t -> ltac_constructor -> raw_tacexpr list -> raw_tacexpr -val of_anti : ?loc:Loc.t -> (?loc:Loc.t -> 'a -> raw_tacexpr) -> 'a or_anti -> raw_tacexpr +val of_anti : ?loc:Loc.t -> ('a -> raw_tacexpr) -> 'a or_anti -> raw_tacexpr -val of_int : ?loc:Loc.t -> int -> raw_tacexpr +val of_int : int located -> raw_tacexpr val of_pair : ?loc:Loc.t -> raw_tacexpr * raw_tacexpr -> raw_tacexpr -val of_variable : ?loc:Loc.t -> Id.t -> raw_tacexpr +val of_variable : Id.t located -> raw_tacexpr -val of_ident : ?loc:Loc.t -> Id.t -> raw_tacexpr +val of_ident : Id.t located -> raw_tacexpr -val of_constr : ?loc:Loc.t -> Constrexpr.constr_expr -> raw_tacexpr +val of_constr : Constrexpr.constr_expr -> raw_tacexpr -val of_open_constr : ?loc:Loc.t -> Constrexpr.constr_expr -> raw_tacexpr +val of_open_constr : Constrexpr.constr_expr -> raw_tacexpr val of_list : ?loc:Loc.t -> raw_tacexpr list -> raw_tacexpr -val of_bindings : ?loc:Loc.t -> bindings -> raw_tacexpr +val of_bindings : bindings -> raw_tacexpr -val of_intro_pattern : ?loc:Loc.t -> intro_pattern -> raw_tacexpr +val of_intro_pattern : intro_pattern -> raw_tacexpr -val of_intro_patterns : ?loc:Loc.t -> intro_pattern list -> raw_tacexpr +val of_intro_patterns : intro_pattern list located -> raw_tacexpr -val of_induction_clause : ?loc:Loc.t -> induction_clause -> raw_tacexpr +val of_induction_clause : induction_clause -> raw_tacexpr -val of_hyp : ?loc:Loc.t -> Id.t -> raw_tacexpr +val of_hyp : ?loc:Loc.t -> Id.t located -> raw_tacexpr (** id ↦ 'Control.hyp @id' *) -val of_exact_hyp : ?loc:Loc.t -> Id.t -> raw_tacexpr +val of_exact_hyp : ?loc:Loc.t -> Id.t located -> raw_tacexpr (** id ↦ 'Control.refine (fun () => Control.hyp @id') *) -- cgit v1.2.3 From de88ba86e9d2a77883365503759eaec96928e9c4 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 4 Aug 2017 17:21:42 +0200 Subject: Introducing quotations for the rewrite tactic. --- src/g_ltac2.ml4 | 36 +++++++++++++++++++++++++++++++++++- src/tac2core.ml | 31 +++++++++---------------------- src/tac2entries.ml | 1 + src/tac2entries.mli | 1 + src/tac2qexpr.mli | 20 +++++++++++++++++++- src/tac2quote.ml | 35 +++++++++++++++++++++++++++++++---- src/tac2quote.mli | 2 ++ theories/Std.v | 15 +++++++++++++++ 8 files changed, 113 insertions(+), 28 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index b10f8d66bd..48a593df28 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -308,7 +308,8 @@ open Tac2entries.Pltac let loc_of_ne_list l = Loc.merge_opt (fst (List.hd l)) (fst (List.last l)) GEXTEND Gram - GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause; + GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause + q_rewriting; anti: [ [ "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] ; @@ -503,6 +504,39 @@ GEXTEND Gram q_induction_clause: [ [ cl = induction_clause -> Tac2quote.of_induction_clause cl ] ] ; + orient: + [ [ "->" -> Loc.tag ~loc:!@loc (Some true) + | "<-" -> Loc.tag ~loc:!@loc (Some false) + | -> Loc.tag ~loc:!@loc None + ]] + ; + rewriter: + [ [ "!"; c = constr_with_bindings -> + (Loc.tag ~loc:!@loc @@ QRepeatPlus,c) + | ["?"| LEFTQMARK]; c = constr_with_bindings -> + (Loc.tag ~loc:!@loc @@ QRepeatStar,c) + | n = lnatural; "!"; c = constr_with_bindings -> + (Loc.tag ~loc:!@loc @@ QPrecisely n,c) + | n = lnatural; ["?" | LEFTQMARK]; c = constr_with_bindings -> + (Loc.tag ~loc:!@loc @@ QUpTo n,c) + | n = lnatural; c = constr_with_bindings -> + (Loc.tag ~loc:!@loc @@ QPrecisely n,c) + | c = constr_with_bindings -> + (Loc.tag ~loc:!@loc @@ QPrecisely (Loc.tag 1), c) + ] ] + ; + oriented_rewriter: + [ [ b = orient; (m, c) = rewriter -> + Loc.tag ~loc:!@loc @@ { + rew_orient = b; + rew_repeat = m; + rew_equatn = c; + } + ] ] + ; + q_rewriting: + [ [ r = oriented_rewriter -> Tac2quote.of_rewriting r ] ] + ; END (** Extension of constr syntax *) diff --git a/src/tac2core.ml b/src/tac2core.ml index 45fa52ff9b..72b4dbfe97 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -883,29 +883,16 @@ let () = add_scope "bindings" begin function | _ -> scope_fail () end -let () = add_scope "intropattern" begin function -| [] -> - let scope = Extend.Aentry Tac2entries.Pltac.q_intropattern in - let act tac = tac in - Tac2entries.ScopeRule (scope, act) -| _ -> scope_fail () -end - -let () = add_scope "intropatterns" begin function -| [] -> - let scope = Extend.Aentry Tac2entries.Pltac.q_intropatterns in - let act tac = tac in - Tac2entries.ScopeRule (scope, act) -| _ -> scope_fail () -end +let add_expr_scope name entry = + add_scope name begin function + | [] -> Tac2entries.ScopeRule (Extend.Aentry entry, (fun e -> e)) + | _ -> scope_fail () + end -let () = add_scope "induction_clause" begin function -| [] -> - let scope = Extend.Aentry Tac2entries.Pltac.q_induction_clause in - let act tac = tac in - Tac2entries.ScopeRule (scope, act) -| _ -> scope_fail () -end +let () = add_expr_scope "intropattern" Tac2entries.Pltac.q_intropattern +let () = add_expr_scope "intropatterns" Tac2entries.Pltac.q_intropatterns +let () = add_expr_scope "induction_clause" Tac2entries.Pltac.q_induction_clause +let () = add_expr_scope "rewriting" Tac2entries.Pltac.q_rewriting let () = add_generic_scope "constr" Pcoq.Constr.constr Stdarg.wit_constr let () = add_generic_scope "open_constr" Pcoq.Constr.constr Stdarg.wit_open_constr diff --git a/src/tac2entries.ml b/src/tac2entries.ml index ce86e8aa33..3aa1ee23b7 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -29,6 +29,7 @@ let q_bindings = Pcoq.Gram.entry_create "tactic:q_bindings" let q_intropattern = Pcoq.Gram.entry_create "tactic:q_intropattern" let q_intropatterns = Pcoq.Gram.entry_create "tactic:q_intropatterns" let q_induction_clause = Pcoq.Gram.entry_create "tactic:q_induction_clause" +let q_rewriting = Pcoq.Gram.entry_create "tactic:q_rewriting" end (** Tactic definition *) diff --git a/src/tac2entries.mli b/src/tac2entries.mli index 1567551246..f5c5a479c4 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -62,4 +62,5 @@ val q_bindings : raw_tacexpr Pcoq.Gram.entry val q_intropattern : raw_tacexpr Pcoq.Gram.entry val q_intropatterns : raw_tacexpr Pcoq.Gram.entry val q_induction_clause : raw_tacexpr Pcoq.Gram.entry +val q_rewriting : raw_tacexpr Pcoq.Gram.entry end diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index 0eb9e9f4b5..f6b8c2c19b 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -63,8 +63,10 @@ type hyp_location = (occurrences * Id.t located or_anti) * Locus.hyp_location_fl type clause = { q_onhyps : hyp_location list option; q_concl_occs : occurrences; } +type constr_with_bindings = (Constrexpr.constr_expr * bindings) located + type destruction_arg = -| QElimOnConstr of (Constrexpr.constr_expr * bindings) located +| QElimOnConstr of constr_with_bindings | QElimOnIdent of Id.t located | QElimOnAnonHyp of int located @@ -76,3 +78,19 @@ type induction_clause_r = { } type induction_clause = induction_clause_r located + +type multi_r = +| QPrecisely of int located +| QUpTo of int located +| QRepeatStar +| QRepeatPlus + +type multi = multi_r located + +type rewriting_r = { + rew_orient : bool option located; + rew_repeat : multi; + rew_equatn : constr_with_bindings; +} + +type rewriting = rewriting_r located diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 1cba488768..7d9c01f3f0 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -109,6 +109,11 @@ let of_bindings (loc, b) = match b with let tl = List.map map tl in std_constructor ?loc "ExplicitBindings" [of_list ?loc tl] +let of_constr_with_bindings (loc, (c, bnd)) = + let c = of_open_constr c in + let bnd = of_bindings bnd in + of_pair ?loc (c, bnd) + let rec of_intro_pattern (loc, pat) = match pat with | QIntroForthcoming b -> std_constructor ?loc "IntroForthcoming" [of_bool b] @@ -179,10 +184,8 @@ let of_clause ?loc cl = ]) let of_destruction_arg ?loc = function -| QElimOnConstr (loc, (c, bnd)) -> - let c = of_open_constr c in - let bnd = of_bindings bnd in - let arg = thunk (of_pair ?loc (c, bnd)) in +| QElimOnConstr c -> + let arg = thunk (of_constr_with_bindings c) in std_constructor ?loc "ElimOnConstr" [arg] | QElimOnIdent id -> std_constructor ?loc "ElimOnIdent" [of_ident id] | QElimOnAnonHyp n -> std_constructor ?loc "ElimOnAnonHyp" [of_int n] @@ -200,6 +203,30 @@ let of_induction_clause (loc, cl) = std_proj "indcl_in", in_; ]) +let of_repeat (loc, r) = match r with +| QPrecisely n -> std_constructor ?loc "Precisely" [of_int n] +| QUpTo n -> std_constructor ?loc "UpTo" [of_int n] +| QRepeatStar -> std_constructor ?loc "RepeatStar" [] +| QRepeatPlus -> std_constructor ?loc "RepeatPlus" [] + +let of_orient loc b = + if b then std_constructor ?loc "LTR" [] + else std_constructor ?loc "RTL" [] + +let of_rewriting (loc, rew) = + let orient = + let (loc, orient) = rew.rew_orient in + of_option ?loc (Option.map (fun b -> of_orient loc b) orient) + in + let repeat = of_repeat rew.rew_repeat in + let equatn = thunk (of_constr_with_bindings rew.rew_equatn) in + let loc = Option.default dummy_loc loc in + CTacRec (loc, [ + std_proj "rew_orient", orient; + std_proj "rew_repeat", repeat; + std_proj "rew_equatn", equatn; + ]) + let of_hyp ?loc id = let loc = Option.default dummy_loc loc in let hyp = CTacRef (AbsKn (control_core "hyp")) in diff --git a/src/tac2quote.mli b/src/tac2quote.mli index 404e9378e0..cb2e406571 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -43,6 +43,8 @@ val of_intro_patterns : intro_pattern list located -> raw_tacexpr val of_induction_clause : induction_clause -> raw_tacexpr +val of_rewriting : rewriting -> raw_tacexpr + val of_hyp : ?loc:Loc.t -> Id.t located -> raw_tacexpr (** id ↦ 'Control.hyp @id' *) diff --git a/theories/Std.v b/theories/Std.v index 19bdc4c82a..09cb3ca0c2 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -86,6 +86,21 @@ Ltac2 Type induction_clause := { indcl_in : clause option; }. +Ltac2 Type repeat := [ +| Precisely (int) +| UpTo (int) +| RepeatStar +| RepeatPlus +]. + +Ltac2 Type orientation := [ LTR | RTL ]. + +Ltac2 Type rewriting := { + rew_orient : orientation option; + rew_repeat : repeat; + rew_equatn : (unit -> constr_with_bindings); +}. + Ltac2 Type evar_flag := bool. Ltac2 Type advanced_flag := bool. -- cgit v1.2.3 From 1f2de88e09c7bb1c0aa111db0d7d50b83f8a62d4 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 4 Aug 2017 18:02:57 +0200 Subject: Exporting the rewrite tactic. --- src/g_ltac2.ml4 | 14 ++++++++------ src/tac2core.ml | 1 + src/tac2entries.ml | 1 + src/tac2entries.mli | 1 + src/tac2quote.mli | 2 ++ src/tac2stdlib.ml | 24 ++++++++++++++++++++++++ src/tac2tactics.ml | 13 +++++++++++++ src/tac2tactics.mli | 13 +++++++++++-- tests/example2.v | 21 +++++++++++++++++++++ theories/Notations.v | 25 +++++++++++++++++++++++++ theories/Std.v | 2 ++ 11 files changed, 109 insertions(+), 8 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 48a593df28..8b373647f3 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -309,7 +309,7 @@ let loc_of_ne_list l = Loc.merge_opt (fst (List.hd l)) (fst (List.last l)) GEXTEND Gram GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause - q_rewriting; + q_rewriting q_clause; anti: [ [ "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] ; @@ -479,12 +479,14 @@ GEXTEND Gram { q_onhyps = Some hl; q_concl_occs = QNoOccurrences } ] ] ; - opt_clause: - [ [ "in"; cl = in_clause -> Some cl - | "at"; occs = occs_nums -> Some { q_onhyps = Some []; q_concl_occs = occs } - | -> None + clause: + [ [ "in"; cl = in_clause -> cl + | "at"; occs = occs_nums -> { q_onhyps = Some []; q_concl_occs = occs } ] ] ; + q_clause: + [ [ cl = clause -> Tac2quote.of_clause cl ] ] + ; concl_occ: [ [ "*"; occs = occs -> occs | -> QNoOccurrences @@ -492,7 +494,7 @@ GEXTEND Gram ; induction_clause: [ [ c = destruction_arg; pat = as_or_and_ipat; eq = eqn_ipat; - cl = opt_clause -> + cl = OPT clause -> Loc.tag ~loc:!@loc @@ { indcl_arg = c; indcl_eqn = eq; diff --git a/src/tac2core.ml b/src/tac2core.ml index 72b4dbfe97..7539e1b697 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -893,6 +893,7 @@ let () = add_expr_scope "intropattern" Tac2entries.Pltac.q_intropattern let () = add_expr_scope "intropatterns" Tac2entries.Pltac.q_intropatterns let () = add_expr_scope "induction_clause" Tac2entries.Pltac.q_induction_clause let () = add_expr_scope "rewriting" Tac2entries.Pltac.q_rewriting +let () = add_expr_scope "clause" Tac2entries.Pltac.q_clause let () = add_generic_scope "constr" Pcoq.Constr.constr Stdarg.wit_constr let () = add_generic_scope "open_constr" Pcoq.Constr.constr Stdarg.wit_open_constr diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 3aa1ee23b7..40d8ff078e 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -30,6 +30,7 @@ let q_intropattern = Pcoq.Gram.entry_create "tactic:q_intropattern" let q_intropatterns = Pcoq.Gram.entry_create "tactic:q_intropatterns" let q_induction_clause = Pcoq.Gram.entry_create "tactic:q_induction_clause" let q_rewriting = Pcoq.Gram.entry_create "tactic:q_rewriting" +let q_clause = Pcoq.Gram.entry_create "tactic:q_clause" end (** Tactic definition *) diff --git a/src/tac2entries.mli b/src/tac2entries.mli index f5c5a479c4..1fe13cda17 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -63,4 +63,5 @@ val q_intropattern : raw_tacexpr Pcoq.Gram.entry val q_intropatterns : raw_tacexpr Pcoq.Gram.entry val q_induction_clause : raw_tacexpr Pcoq.Gram.entry val q_rewriting : raw_tacexpr Pcoq.Gram.entry +val q_clause : raw_tacexpr Pcoq.Gram.entry end diff --git a/src/tac2quote.mli b/src/tac2quote.mli index cb2e406571..ddb39326d1 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -41,6 +41,8 @@ val of_intro_pattern : intro_pattern -> raw_tacexpr val of_intro_patterns : intro_pattern list located -> raw_tacexpr +val of_clause : ?loc:Loc.t -> clause -> raw_tacexpr + val of_induction_clause : induction_clause -> raw_tacexpr val of_rewriting : rewriting -> raw_tacexpr diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 7e421c8577..e8e63f520c 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -141,6 +141,21 @@ let to_induction_clause = function | _ -> assert false +let to_multi = function +| ValBlk (0, [| n |]) -> Precisely (Value.to_int n) +| ValBlk (1, [| n |]) -> UpTo (Value.to_int n) +| ValInt 0 -> RepeatStar +| ValInt 1 -> RepeatPlus +| _ -> assert false + +let to_rewriting = function +| ValBlk (0, [| orient; repeat; c |]) -> + let orient = Value.to_option Value.to_bool orient in + let repeat = to_multi repeat in + let c = thaw c >>= fun c -> return (to_constr_with_bindings c) in + (orient, repeat, c) +| _ -> assert false + (** Standard tactics sharing their implementation with Ltac1 *) let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } @@ -304,6 +319,15 @@ let () = define_prim2 "tac_lazy" begin fun flags cl -> Tactics.reduce (Lazy flags) cl end +let () = define_prim4 "tac_rewrite" begin fun ev rw cl by -> + let ev = Value.to_bool ev in + let rw = Value.to_list to_rewriting rw in + let cl = to_clause cl in + let to_tac t = Proofview.tclIGNORE (thaw t) in + let by = Value.to_option to_tac by in + Tac2tactics.rewrite ev rw cl by +end + (** Tactics from coretactics *) let () = define_prim0 "tac_reflexivity" Tactics.intros_reflexivity diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index 439250db78..e7e15544af 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -44,3 +44,16 @@ let map_induction_clause ((clear, arg), eqn, as_, occ) = let induction_destruct isrec ev ic using = let ic = List.map map_induction_clause ic in Tactics.induction_destruct isrec ev (ic, using) + +type rewriting = + bool option * + multi * + EConstr.constr with_bindings tactic + +let rewrite ev rw cl by = + let map_rw (orient, repeat, c) = + (Option.default true orient, repeat, None, delayed_of_tactic c) + in + let rw = List.map map_rw rw in + let by = Option.map (fun tac -> Tacticals.New.tclCOMPLETE tac, Equality.Naive) by in + Equality.general_multi_rewrite ev rw cl by diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli index f29793411a..93cc6ecd68 100644 --- a/src/tac2tactics.mli +++ b/src/tac2tactics.mli @@ -7,6 +7,7 @@ (************************************************************************) open Names +open Locus open Misctypes open Tactypes open Proofview @@ -21,7 +22,15 @@ type induction_clause = EConstr.constr with_bindings tactic destruction_arg * intro_pattern_naming option * or_and_intro_pattern option * - Locus.clause option + clause option val induction_destruct : rec_flag -> evars_flag -> - induction_clause list -> EConstr.constr with_bindings option -> unit Proofview.tactic + induction_clause list -> EConstr.constr with_bindings option -> unit tactic + +type rewriting = + bool option * + multi * + EConstr.constr with_bindings tactic + +val rewrite : + evars_flag -> rewriting list -> clause -> unit tactic option -> unit tactic diff --git a/tests/example2.v b/tests/example2.v index 812f9172c9..526cbc39f5 100644 --- a/tests/example2.v +++ b/tests/example2.v @@ -95,3 +95,24 @@ Proof. intros b1 b2. destruct &b1 as [|], &b2 as [|]; split. Qed. + +Goal forall n m, n = 0 -> n + m = m. +Proof. +intros n m Hn. +rewrite &Hn; split. +Qed. + +Goal forall n m p, n = m -> p = m -> 0 = n -> p = 0. +Proof. +intros n m p He He' Hn. +rewrite &He, <- &He' in Hn. +rewrite &Hn. +split. +Qed. + +Goal forall n m, (m = n -> n = m) -> m = n -> n = 0 -> m = 0. +Proof. +intros n m He He' He''. +rewrite <- &He by Std.assumption (). +Control.refine (fun () => &He''). +Qed. diff --git a/theories/Notations.v b/theories/Notations.v index 20f01c3b48..4ce9fc0dbd 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -138,3 +138,28 @@ Ltac2 Notation "edestruct" ic(list1(induction_clause, ",")) use(thunk(opt(seq("using", constr, bindings)))) := destruct0 true ic use. + +Ltac2 rewrite0 ev rw cl tac := + let tac := match tac with + | None => None + | Some p => + let ((_, tac)) := p in + Some tac + end in + let cl := match cl with + | None => { Std.on_hyps := Some []; Std.on_concl := Std.AllOccurrences } + | Some cl => cl + end in + Std.rewrite ev rw cl tac. + +Ltac2 Notation "rewrite" + rw(list1(rewriting, ",")) + cl(opt(clause)) + tac(opt(seq("by", thunk(tactic)))) := + rewrite0 false rw cl tac. + +Ltac2 Notation "erewrite" + rw(list1(rewriting, ",")) + cl(opt(clause)) + tac(opt(seq("by", thunk(tactic)))) := + rewrite0 true rw cl tac. diff --git a/theories/Std.v b/theories/Std.v index 09cb3ca0c2..695ea26444 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -134,6 +134,8 @@ Ltac2 @ external cbv : red_flags -> clause -> unit := "ltac2" "tac_cbv". Ltac2 @ external cbn : red_flags -> clause -> unit := "ltac2" "tac_cbn". Ltac2 @ external lazy : red_flags -> clause -> unit := "ltac2" "tac_lazy". +Ltac2 @ external rewrite : evar_flag -> rewriting list -> clause -> (unit -> unit) option -> unit := "ltac2" "tac_rewrite". + Ltac2 @ external reflexivity : unit -> unit := "ltac2" "tac_reflexivity". Ltac2 @ external assumption : unit -> unit := "ltac2" "tac_assumption". -- cgit v1.2.3 From dbaf8dd6b150619ac04b33ae4d581432cb5cefe0 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 5 Aug 2017 14:46:18 +0200 Subject: More notations for basic tactics. --- tests/example2.v | 7 +++++++ theories/Notations.v | 17 +++++++++++++---- 2 files changed, 20 insertions(+), 4 deletions(-) diff --git a/tests/example2.v b/tests/example2.v index 526cbc39f5..76f069a5ae 100644 --- a/tests/example2.v +++ b/tests/example2.v @@ -116,3 +116,10 @@ intros n m He He' He''. rewrite <- &He by Std.assumption (). Control.refine (fun () => &He''). Qed. + +Goal forall n (r := if true then n else 0), r = n. +Proof. +intros n r. +hnf in r. +split. +Qed. diff --git a/theories/Notations.v b/theories/Notations.v index 4ce9fc0dbd..4bba9a7495 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -139,6 +139,18 @@ Ltac2 Notation "edestruct" use(thunk(opt(seq("using", constr, bindings)))) := destruct0 true ic use. +Ltac2 default_on_concl cl := +match cl with +| None => { Std.on_hyps := Some []; Std.on_concl := Std.AllOccurrences } +| Some cl => cl +end. + +Ltac2 Notation "red" cl(opt(clause)) := + Std.red (default_on_concl cl). + +Ltac2 Notation "hnf" cl(opt(clause)) := + Std.hnf (default_on_concl cl). + Ltac2 rewrite0 ev rw cl tac := let tac := match tac with | None => None @@ -146,10 +158,7 @@ Ltac2 rewrite0 ev rw cl tac := let ((_, tac)) := p in Some tac end in - let cl := match cl with - | None => { Std.on_hyps := Some []; Std.on_concl := Std.AllOccurrences } - | Some cl => cl - end in + let cl := default_on_concl cl in Std.rewrite ev rw cl tac. Ltac2 Notation "rewrite" -- cgit v1.2.3 From 6e6f348958cc333040991ca3dc2525a7c91dc9c0 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 5 Aug 2017 15:42:29 +0200 Subject: Exporting more reduction functions. --- src/tac2stdlib.ml | 48 ++++++++++++++++++++++++++++++++++++++++++++++++ src/tac2tactics.ml | 25 +++++++++++++++++++++++++ src/tac2tactics.mli | 8 ++++++++ theories/Std.v | 6 ++++++ 4 files changed, 87 insertions(+) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index e8e63f520c..3cfd0b5626 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -90,6 +90,13 @@ let to_red_flag = function } | _ -> assert false +let to_pattern_with_occs pat = + to_pair Value.to_pattern (fun occ -> to_occurrences to_int_or_var occ) pat + +let to_constr_with_occs c = + let (c, occ) = to_pair Value.to_constr (fun occ -> to_occurrences to_int_or_var occ) c in + (occ, c) + let rec to_intro_pattern = function | ValBlk (0, [| b |]) -> IntroForthcoming (Value.to_bool b) | ValBlk (1, [| pat |]) -> IntroNaming (to_intro_pattern_naming pat) @@ -301,6 +308,13 @@ let () = define_prim1 "tac_hnf" begin fun cl -> Tactics.reduce Hnf cl end +let () = define_prim3 "tac_simpl" begin fun flags where cl -> + let flags = to_red_flag flags in + let where = Value.to_option to_pattern_with_occs where in + let cl = to_clause cl in + Tac2tactics.simpl flags where cl +end + let () = define_prim2 "tac_cbv" begin fun flags cl -> let flags = to_red_flag flags in let cl = to_clause cl in @@ -319,6 +333,40 @@ let () = define_prim2 "tac_lazy" begin fun flags cl -> Tactics.reduce (Lazy flags) cl end +let () = define_prim2 "tac_unfold" begin fun refs cl -> + let map v = + let (ref, occ) = to_pair to_evaluable_ref (fun occ -> to_occurrences to_int_or_var occ) v in + (occ, ref) + in + let refs = Value.to_list map refs in + let cl = to_clause cl in + Tactics.reduce (Unfold refs) cl +end + +let () = define_prim2 "tac_fold" begin fun args cl -> + let args = Value.to_list Value.to_constr args in + let cl = to_clause cl in + Tactics.reduce (Fold args) cl +end + +let () = define_prim2 "tac_pattern" begin fun where cl -> + let where = Value.to_list to_constr_with_occs where in + let cl = to_clause cl in + Tactics.reduce (Pattern where) cl +end + +let () = define_prim2 "tac_vm" begin fun where cl -> + let where = Value.to_option to_pattern_with_occs where in + let cl = to_clause cl in + Tac2tactics.vm where cl +end + +let () = define_prim2 "tac_native" begin fun where cl -> + let where = Value.to_option to_pattern_with_occs where in + let cl = to_clause cl in + Tac2tactics.native where cl +end + let () = define_prim4 "tac_rewrite" begin fun ev rw cl by -> let ev = Value.to_bool ev in let rw = Value.to_list to_rewriting rw in diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index e7e15544af..50c1df922e 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -6,8 +6,12 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open Util +open Names +open Globnames open Misctypes open Tactypes +open Genredexpr open Tactics open Proofview open Tacmach.New @@ -57,3 +61,24 @@ let rewrite ev rw cl by = let rw = List.map map_rw rw in let by = Option.map (fun tac -> Tacticals.New.tclCOMPLETE tac, Equality.Naive) by in Equality.general_multi_rewrite ev rw cl by + +(** Ltac interface treats differently global references than other term + arguments in reduction expressions. In Ltac1, this is done at parsing time. + Instead, we parse indifferently any pattern and dispatch when the tactic is + called. *) +let map_pattern_with_occs (pat, occ) = match pat with +| Pattern.PRef (ConstRef cst) -> (occ, Inl (EvalConstRef cst)) +| Pattern.PRef (VarRef id) -> (occ, Inl (EvalVarRef id)) +| _ -> (occ, Inr pat) + +let simpl flags where cl = + let where = Option.map map_pattern_with_occs where in + Tactics.reduce (Simpl (flags, where)) cl + +let vm where cl = + let where = Option.map map_pattern_with_occs where in + Tactics.reduce (CbvVm where) cl + +let native where cl = + let where = Option.map map_pattern_with_occs where in + Tactics.reduce (CbvNative where) cl diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli index 93cc6ecd68..affbbbbdd7 100644 --- a/src/tac2tactics.mli +++ b/src/tac2tactics.mli @@ -8,6 +8,7 @@ open Names open Locus +open Genredexpr open Misctypes open Tactypes open Proofview @@ -34,3 +35,10 @@ type rewriting = val rewrite : evars_flag -> rewriting list -> clause -> unit tactic option -> unit tactic + +val simpl : evaluable_global_reference glob_red_flag -> + (Pattern.constr_pattern * occurrences_expr) option -> clause -> unit tactic + +val vm : (Pattern.constr_pattern * occurrences_expr) option -> clause -> unit tactic + +val native : (Pattern.constr_pattern * occurrences_expr) option -> clause -> unit tactic diff --git a/theories/Std.v b/theories/Std.v index 695ea26444..43ccb06192 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -130,9 +130,15 @@ Ltac2 @ external induction : evar_flag -> induction_clause list -> Ltac2 @ external red : clause -> unit := "ltac2" "tac_red". Ltac2 @ external hnf : clause -> unit := "ltac2" "tac_hnf". +Ltac2 @ external simpl : red_flags -> (pattern * occurrences) option -> clause -> unit := "ltac2" "tac_simpl". Ltac2 @ external cbv : red_flags -> clause -> unit := "ltac2" "tac_cbv". Ltac2 @ external cbn : red_flags -> clause -> unit := "ltac2" "tac_cbn". Ltac2 @ external lazy : red_flags -> clause -> unit := "ltac2" "tac_lazy". +Ltac2 @ external unfold : (evaluable_reference * occurrences) list -> clause -> unit := "ltac2" "tac_unfold". +Ltac2 @ external fold : constr list -> clause -> unit := "ltac2" "tac_fold". +Ltac2 @ external pattern : (constr * occurrences) list -> clause -> unit := "ltac2" "tac_pattern". +Ltac2 @ external vm : (pattern * occurrences) option -> clause -> unit := "ltac2" "tac_vm". +Ltac2 @ external native : (pattern * occurrences) option -> clause -> unit := "ltac2" "tac_native". Ltac2 @ external rewrite : evar_flag -> rewriting list -> clause -> (unit -> unit) option -> unit := "ltac2" "tac_rewrite". -- cgit v1.2.3 From 77150cc524f5cbdc9bf340be03f31e7f7542c98d Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 5 Aug 2017 17:14:21 +0200 Subject: Introducing grammar-free tactic notations. --- doc/ltac2.md | 32 +++++++++++++++++++++++++++++ src/tac2entries.ml | 60 +++++++++++++++++++++++++++++++++++++++++++++--------- src/tac2env.ml | 47 +++++++++++++++++++++++++++++++++--------- src/tac2env.mli | 13 ++++++++---- src/tac2expr.mli | 7 ++++++- src/tac2intern.ml | 30 ++++++++++++++++++++++++--- src/tac2print.ml | 2 +- src/tac2quote.ml | 4 ++-- 8 files changed, 164 insertions(+), 31 deletions(-) diff --git a/doc/ltac2.md b/doc/ltac2.md index c2d930c9b6..b3596b2977 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -654,6 +654,38 @@ Beware that the order of evaluation of multiple let-bindings is not specified, so that you may have to resort to thunking to ensure that side-effects are performed at the right time. +## Abbreviations + +There exists a special kind of notations, called abbreviations, that is designed +so that it does not add any parsing rules. It is similar in spirit to Coq +abbreviations, insofar as its main purpose is to give an absolute name to a +piece of pure syntax, which can be transparently referred by this name as if it +were a proper definition. Abbreviations are introduced by the following +syntax. + +``` +VERNAC ::= +| "Ltac2" "Notation" IDENT ":=" TERM +``` + +The abbreviation can then be manipulated just as a normal Ltac2 definition, +except that it is expanded at internalization time into the given expression. +Furthermore, in order to make this kind of construction useful in practice in +an effectful language such as Ltac2, any syntactic argument to an abbreviation +is thunked on-the-fly during its expansion. + +For instance, suppose that we define the following. +``` +Ltac2 Notation foo := fun x => x (). +``` +Then we have the following expansion at internalization time. +``` +foo 0 ↦ (fun x => x ()) (fun _ => 0) +``` + +Note that abbreviations are not typechecked at all, and may result in typing +errors after expansion. + # TODO - Implement deep pattern-matching. diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 40d8ff078e..1dd8410d2a 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -42,14 +42,14 @@ type tacdef = { } let perform_tacdef visibility ((sp, kn), def) = - let () = if not def.tacdef_local then Tac2env.push_ltac visibility sp kn in + let () = if not def.tacdef_local then Tac2env.push_ltac visibility sp (TacConstant kn) in Tac2env.define_global kn (def.tacdef_expr, def.tacdef_type) let load_tacdef i obj = perform_tacdef (Until i) obj let open_tacdef i obj = perform_tacdef (Exactly i) obj let cache_tacdef ((sp, kn), def) = - let () = Tac2env.push_ltac (Until 1) sp kn in + let () = Tac2env.push_ltac (Until 1) sp (TacConstant kn) in Tac2env.define_global kn (def.tacdef_expr, def.tacdef_type) let subst_tacdef (subst, def) = @@ -599,7 +599,43 @@ let inTac2Notation : synext -> obj = subst_function = subst_synext; classify_function = classify_synext} -let register_notation ?(local = false) tkn lev body = +type abbreviation = { + abbr_body : raw_tacexpr; +} + +let perform_abbreviation visibility ((sp, kn), abbr) = + let () = Tac2env.push_ltac visibility sp (TacAlias kn) in + Tac2env.define_alias kn abbr.abbr_body + +let load_abbreviation i obj = perform_abbreviation (Until i) obj +let open_abbreviation i obj = perform_abbreviation (Exactly i) obj + +let cache_abbreviation ((sp, kn), abbr) = + let () = Tac2env.push_ltac (Until 1) sp (TacAlias kn) in + Tac2env.define_alias kn abbr.abbr_body + +let subst_abbreviation (subst, abbr) = + let body' = subst_rawexpr subst abbr.abbr_body in + if body' == abbr.abbr_body then abbr + else { abbr_body = body' } + +let classify_abbreviation o = Substitute o + +let inTac2Abbreviation : abbreviation -> obj = + declare_object {(default_object "TAC2-ABBREVIATION") with + cache_function = cache_abbreviation; + load_function = load_abbreviation; + open_function = open_abbreviation; + subst_function = subst_abbreviation; + classify_function = classify_abbreviation} + +let register_notation ?(local = false) tkn lev body = match tkn, lev with +| [SexprRec (_, (_, Some id), [])], None -> + (** Tactic abbreviation *) + let body = Tac2intern.globalize Id.Set.empty body in + let abbr = { abbr_body = body } in + ignore (Lib.add_leaf id (inTac2Abbreviation abbr)) +| _ -> (** Check that the tokens make sense *) let entries = List.map ParseToken.parse_token tkn in let fold accu tok = match tok with @@ -642,14 +678,18 @@ let print_ltac ref = try Tac2env.locate_ltac qid with Not_found -> user_err ?loc (str "Unknown tactic " ++ pr_qualid qid) in - let (e, _, (_, t)) = Tac2env.interp_global kn in - let name = int_name () in - Feedback.msg_notice ( - hov 0 ( - hov 2 (pr_qualid qid ++ spc () ++ str ":" ++ spc () ++ pr_glbtype name t) ++ fnl () ++ - hov 2 (pr_qualid qid ++ spc () ++ str ":=" ++ spc () ++ pr_glbexpr e) + match kn with + | TacConstant kn -> + let (e, _, (_, t)) = Tac2env.interp_global kn in + let name = int_name () in + Feedback.msg_notice ( + hov 0 ( + hov 2 (pr_qualid qid ++ spc () ++ str ":" ++ spc () ++ pr_glbtype name t) ++ fnl () ++ + hov 2 (pr_qualid qid ++ spc () ++ str ":=" ++ spc () ++ pr_glbexpr e) + ) ) - ) + | TacAlias kn -> + Feedback.msg_notice (str "Alias to ...") (** Calling tactics *) diff --git a/src/tac2env.ml b/src/tac2env.ml index a75500eae7..ac2bd5fc23 100644 --- a/src/tac2env.ml +++ b/src/tac2env.ml @@ -32,6 +32,7 @@ type ltac_state = { ltac_constructors : constructor_data KNmap.t; ltac_projections : projection_data KNmap.t; ltac_types : glb_quant_typedef KNmap.t; + ltac_aliases : raw_tacexpr KNmap.t; } let empty_state = { @@ -39,6 +40,7 @@ let empty_state = { ltac_constructors = KNmap.empty; ltac_projections = KNmap.empty; ltac_types = KNmap.empty; + ltac_aliases = KNmap.empty; } let ltac_state = Summary.ref empty_state ~name:"ltac2-state" @@ -87,6 +89,12 @@ let define_type kn e = let interp_type kn = KNmap.find kn ltac_state.contents.ltac_types +let define_alias kn tac = + let state = !ltac_state in + ltac_state := { state with ltac_aliases = KNmap.add kn tac state.ltac_aliases } + +let interp_alias kn = KNmap.find kn ltac_state.contents.ltac_aliases + module ML = struct type t = ml_tactic_name @@ -115,11 +123,30 @@ struct id, (DirPath.repr dir) end +type tacref = Tac2expr.tacref = +| TacConstant of ltac_constant +| TacAlias of ltac_alias + +module TacRef = +struct +type t = tacref +let compare r1 r2 = match r1, r2 with +| TacConstant c1, TacConstant c2 -> KerName.compare c1 c2 +| TacAlias c1, TacAlias c2 -> KerName.compare c1 c2 +| TacConstant _, TacAlias _ -> -1 +| TacAlias _, TacConstant _ -> 1 + +let equal r1 r2 = compare r1 r2 == 0 + +end + module KnTab = Nametab.Make(FullPath)(KerName) +module RfTab = Nametab.Make(FullPath)(TacRef) +module RfMap = Map.Make(TacRef) type nametab = { - tab_ltac : KnTab.t; - tab_ltac_rev : full_path KNmap.t; + tab_ltac : RfTab.t; + tab_ltac_rev : full_path RfMap.t; tab_cstr : KnTab.t; tab_cstr_rev : full_path KNmap.t; tab_type : KnTab.t; @@ -129,8 +156,8 @@ type nametab = { } let empty_nametab = { - tab_ltac = KnTab.empty; - tab_ltac_rev = KNmap.empty; + tab_ltac = RfTab.empty; + tab_ltac_rev = RfMap.empty; tab_cstr = KnTab.empty; tab_cstr_rev = KNmap.empty; tab_type = KnTab.empty; @@ -143,22 +170,22 @@ let nametab = Summary.ref empty_nametab ~name:"ltac2-nametab" let push_ltac vis sp kn = let tab = !nametab in - let tab_ltac = KnTab.push vis sp kn tab.tab_ltac in - let tab_ltac_rev = KNmap.add kn sp tab.tab_ltac_rev in + let tab_ltac = RfTab.push vis sp kn tab.tab_ltac in + let tab_ltac_rev = RfMap.add kn sp tab.tab_ltac_rev in nametab := { tab with tab_ltac; tab_ltac_rev } let locate_ltac qid = let tab = !nametab in - KnTab.locate qid tab.tab_ltac + RfTab.locate qid tab.tab_ltac let locate_extended_all_ltac qid = let tab = !nametab in - KnTab.find_prefixes qid tab.tab_ltac + RfTab.find_prefixes qid tab.tab_ltac let shortest_qualid_of_ltac kn = let tab = !nametab in - let sp = KNmap.find kn tab.tab_ltac_rev in - KnTab.shortest_qualid Id.Set.empty sp tab.tab_ltac + let sp = RfMap.find kn tab.tab_ltac_rev in + RfTab.shortest_qualid Id.Set.empty sp tab.tab_ltac let push_constructor vis sp kn = let tab = !nametab in diff --git a/src/tac2env.mli b/src/tac2env.mli index fea03c4285..e4cc8387c5 100644 --- a/src/tac2env.mli +++ b/src/tac2env.mli @@ -61,12 +61,17 @@ type projection_data = { val define_projection : ltac_projection -> projection_data -> unit val interp_projection : ltac_projection -> projection_data +(** {5 Toplevel definition of aliases} *) + +val define_alias : ltac_constant -> raw_tacexpr -> unit +val interp_alias : ltac_constant -> raw_tacexpr + (** {5 Name management} *) -val push_ltac : visibility -> full_path -> ltac_constant -> unit -val locate_ltac : qualid -> ltac_constant -val locate_extended_all_ltac : qualid -> ltac_constant list -val shortest_qualid_of_ltac : ltac_constant -> qualid +val push_ltac : visibility -> full_path -> tacref -> unit +val locate_ltac : qualid -> tacref +val locate_extended_all_ltac : qualid -> tacref list +val shortest_qualid_of_ltac : tacref -> qualid val push_constructor : visibility -> full_path -> ltac_constructor -> unit val locate_constructor : qualid -> ltac_constructor diff --git a/src/tac2expr.mli b/src/tac2expr.mli index 10d8c1d421..7efb85cbb0 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -18,10 +18,15 @@ type lid = Id.t type uid = Id.t type ltac_constant = KerName.t +type ltac_alias = KerName.t type ltac_constructor = KerName.t type ltac_projection = KerName.t type type_constant = KerName.t +type tacref = +| TacConstant of ltac_constant +| TacAlias of ltac_alias + type 'a or_relid = | RelId of qualid located | AbsKn of 'a @@ -88,7 +93,7 @@ type raw_patexpr = type raw_tacexpr = | CTacAtm of atom located -| CTacRef of ltac_constant or_relid +| CTacRef of tacref or_relid | CTacCst of Loc.t * ltac_constructor or_tuple or_relid | CTacFun of Loc.t * (raw_patexpr * raw_typexpr option) list * raw_tacexpr | CTacApp of Loc.t * raw_tacexpr * raw_tacexpr list diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 32ed211ad0..2b1dde7553 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -654,6 +654,10 @@ let expand_pattern avoid bnd = let nas = List.rev_map (fun (na, _, _) -> na) bnd in (nas, expand) +let is_alias env qid = match get_variable env qid with +| ArgArg (TacAlias _) -> true +| ArgVar _ | (ArgArg (TacConstant _)) -> false + let rec intern_rec env = function | CTacAtm (_, atm) -> intern_atm env atm | CTacRef qid -> @@ -661,9 +665,12 @@ let rec intern_rec env = function | ArgVar (_, id) -> let sch = Id.Map.find id env.env_var in (GTacVar id, fresh_mix_type_scheme env sch) - | ArgArg kn -> + | ArgArg (TacConstant kn) -> let (_, _, sch) = Tac2env.interp_global kn in (GTacRef kn, fresh_type_scheme env sch) + | ArgArg (TacAlias kn) -> + let e = Tac2env.interp_alias kn in + intern_rec env e end | CTacCst (loc, qid) -> let kn = get_constructor env qid in @@ -682,6 +689,20 @@ let rec intern_rec env = function | CTacApp (loc, CTacCst (_, qid), args) -> let kn = get_constructor env qid in intern_constructor env loc kn args +| CTacApp (loc, CTacRef qid, args) when is_alias env qid -> + let kn = match get_variable env qid with + | ArgArg (TacAlias kn) -> kn + | ArgVar _ | (ArgArg (TacConstant _)) -> assert false + in + let e = Tac2env.interp_alias kn in + let map arg = + (** Thunk alias arguments *) + let loc = loc_of_tacexpr arg in + let var = [CPatVar (Some loc, Anonymous), Some (CTypRef (loc, AbsKn (Tuple 0), []))] in + CTacFun (loc, var, arg) + in + let args = List.map map args in + intern_rec env (CTacApp (loc, e, args)) | CTacApp (loc, f, args) -> let loc = loc_of_tacexpr f in let (f, ft) = intern_rec env f in @@ -1372,9 +1393,12 @@ let rec subst_rawtype subst t = match t with let subst_tacref subst ref = match ref with | RelId _ -> ref -| AbsKn kn -> +| AbsKn (TacConstant kn) -> let kn' = subst_kn subst kn in - if kn' == kn then ref else AbsKn kn' + if kn' == kn then ref else AbsKn (TacConstant kn') +| AbsKn (TacAlias kn) -> + let kn' = subst_kn subst kn in + if kn' == kn then ref else AbsKn (TacAlias kn') let subst_projection subst prj = match prj with | RelId _ -> prj diff --git a/src/tac2print.ml b/src/tac2print.ml index 28f9516f65..b679e030fd 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -143,7 +143,7 @@ let pr_glbexpr_gen lvl c = | GTacAtm atm -> pr_atom atm | GTacVar id -> Id.print id | GTacRef gr -> - let qid = shortest_qualid_of_ltac gr in + let qid = shortest_qualid_of_ltac (TacConstant gr) in Libnames.pr_qualid qid | GTacFun (nas, c) -> let nas = pr_sequence pr_name nas in diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 7d9c01f3f0..3e25e1cadb 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -229,10 +229,10 @@ let of_rewriting (loc, rew) = let of_hyp ?loc id = let loc = Option.default dummy_loc loc in - let hyp = CTacRef (AbsKn (control_core "hyp")) in + let hyp = CTacRef (AbsKn (TacConstant (control_core "hyp"))) in CTacApp (loc, hyp, [of_ident id]) let of_exact_hyp ?loc id = let loc = Option.default dummy_loc loc in - let refine = CTacRef (AbsKn (control_core "refine")) in + let refine = CTacRef (AbsKn (TacConstant (control_core "refine"))) in CTacApp (loc, refine, [thunk (of_hyp ~loc id)]) -- cgit v1.2.3 From e1ea058fb664be58371237e5a6dbe0ec570448d5 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 7 Aug 2017 15:15:36 +0200 Subject: Defining a few base tacticals. --- src/tac2core.ml | 25 +++++++++++++++++++++++ theories/Control.v | 13 ++++++++++++ theories/Notations.v | 56 +++++++++++++++++++++++++++++++++++++++++++++++++++- 3 files changed, 93 insertions(+), 1 deletion(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 7539e1b697..08f61f2c6c 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -574,6 +574,27 @@ let prm_with_holes : ml_tactic = function Tacticals.New.tclWITHHOLES false (interp_app f [ans]) sigma | _ -> assert false +let prm_progress : ml_tactic = function +| [f] -> Proofview.tclPROGRESS (thaw f) +| _ -> assert false + +let prm_abstract : ml_tactic = function +| [id; f] -> + let id = Value.to_option Value.to_ident id in + Tactics.tclABSTRACT id (Proofview.tclIGNORE (thaw f)) >>= fun () -> + return v_unit +| _ -> assert false + +let prm_time : ml_tactic = function +| [s; f] -> + let s = Value.to_option Value.to_string s in + Proofview.tclTIME s (thaw f) +| _ -> assert false + +let prm_check_interrupt : ml_tactic = function +| [_] -> Proofview.tclCHECKINTERRUPT >>= fun () -> return v_unit +| _ -> assert false + (** Registering *) let () = Tac2env.define_primitive (pname "print") prm_print @@ -632,6 +653,10 @@ let () = Tac2env.define_primitive (pname "hyp") prm_hyp let () = Tac2env.define_primitive (pname "hyps") prm_hyps let () = Tac2env.define_primitive (pname "refine") prm_refine let () = Tac2env.define_primitive (pname "with_holes") prm_with_holes +let () = Tac2env.define_primitive (pname "progress") prm_progress +let () = Tac2env.define_primitive (pname "abstract") prm_abstract +let () = Tac2env.define_primitive (pname "time") prm_time +let () = Tac2env.define_primitive (pname "check_interrupt") prm_check_interrupt (** ML types *) diff --git a/theories/Control.v b/theories/Control.v index a8b92aced2..071c2ea8ce 100644 --- a/theories/Control.v +++ b/theories/Control.v @@ -34,6 +34,8 @@ Ltac2 @ external new_goal : evar -> unit := "ltac2" "new_goal". already defined in the current state, don't do anything. Panics if the evar is not in the current state. *) +Ltac2 @ external progress : (unit -> 'a) -> 'a := "ltac2" "progress". + (** Goal inspection *) Ltac2 @ external goal : unit -> constr := "ltac2" "goal". @@ -61,3 +63,14 @@ Ltac2 @ external with_holes : (unit -> 'a) -> ('a -> 'b) -> 'b := "ltac2" "with_ (** [with_holes x f] evaluates [x], then apply [f] to the result, and fails if all evars generated by the call to [x] have not been solved when [f] returns. *) + +(** Misc *) + +Ltac2 @ external time : string option -> (unit -> 'a) -> 'a := "ltac2" "time". +(** Displays the time taken by a tactic to evaluate. *) + +Ltac2 @ external abstract : ident option -> (unit -> unit) -> unit := "ltac2" "abstract". +(** Abstract a subgoal. *) + +Ltac2 @ external check_interrupt : unit -> unit := "ltac2" "check_interrupt". +(** For internal use. *) diff --git a/theories/Notations.v b/theories/Notations.v index 4bba9a7495..f11cfbde6e 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -7,7 +7,61 @@ (************************************************************************) Require Import Ltac2.Init. -Require Ltac2.Control Ltac2.Std. +Require Ltac2.Control Ltac2.Int Ltac2.Std. + +(** Tacticals *) + +Ltac2 orelse t f := +match Control.case t with +| Err e => f e +| Val ans => + let ((x, k)) := ans in + Control.plus (fun _ => x) k +end. + +Ltac2 ifcatch t s f := +match Control.case t with +| Err e => f e +| Val ans => + let ((x, k)) := ans in + Control.plus (fun _ => s x) (fun e => s (k e)) +end. + +Ltac2 try0 t := Control.enter (fun _ => orelse t (fun _ => ())). + +Ltac2 Notation try := try0. + +Ltac2 rec repeat0 (t : unit -> unit) := + Control.enter (fun () => + ifcatch (fun _ => Control.progress t) + (fun _ => Control.check_interrupt (); repeat0 t) (fun _ => ())). + +Ltac2 Notation repeat := repeat0. + +Ltac2 do0 n t := + let rec aux n t := match Int.equal n 0 with + | true => () + | false => t (); aux (Int.sub n 1) t + end in + aux (n ()) t. + +Ltac2 Notation do := do0. + +Ltac2 Notation once := Control.once. + +Ltac2 progress0 tac := Control.enter (fun _ => Control.progress tac). + +Ltac2 Notation progress := progress0. + +Ltac2 time0 tac := Control.time None tac. + +Ltac2 Notation time := time0. + +Ltac2 abstract0 tac := Control.abstract None tac. + +Ltac2 Notation abstract := abstract0. + +(** Base tactics *) (** Enter and check evar resolution *) Ltac2 enter_h ev f arg := -- cgit v1.2.3 From c0f72bb07442075ee1dc66b4902b5c4681d220cf Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 7 Aug 2017 16:22:45 +0200 Subject: Fix parsing of parenthesized expressions. --- src/g_ltac2.ml4 | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 8b373647f3..1cab918ea4 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -119,8 +119,8 @@ GEXTEND Gram | e = SELF; ".("; qid = Prim.qualid; ")" -> CTacPrj (!@loc, e, RelId qid) | e = SELF; ".("; qid = Prim.qualid; ")"; ":="; r = tac2expr LEVEL "5" -> CTacSet (!@loc, e, RelId qid, r) ] | "0" - [ "("; a = tac2expr LEVEL "5"; ")" -> a - | "("; a = tac2expr; ":"; t = tac2type; ")" -> CTacCnv (!@loc, a, t) + [ "("; a = SELF; ")" -> a + | "("; a = SELF; ":"; t = tac2type; ")" -> CTacCnv (!@loc, a, t) | "()" -> CTacCst (!@loc, AbsKn (Tuple 0)) | "("; ")" -> CTacCst (!@loc, AbsKn (Tuple 0)) | "["; a = LIST0 tac2expr LEVEL "5" SEP ";"; "]" -> CTacLst (Loc.tag ~loc:!@loc a) -- cgit v1.2.3 From 49f0f89e3a7905679d758017ccaeba64c0ca79b1 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 7 Aug 2017 16:26:35 +0200 Subject: Fix location of not-unit warning. --- src/tac2intern.ml | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 2b1dde7553..bf7e93cb9e 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -750,9 +750,10 @@ let rec intern_rec env = function let () = unify ~loc env t tc in (e, tc) | CTacSeq (loc, e1, e2) -> + let loc1 = loc_of_tacexpr e1 in let (e1, t1) = intern_rec env e1 in let (e2, t2) = intern_rec env e2 in - let () = check_elt_unit loc env t1 in + let () = check_elt_unit loc1 env t1 in (GTacLet (false, [Anonymous, e1], e2), t2) | CTacCse (loc, e, pl) -> intern_case env loc e pl -- cgit v1.2.3 From bdd7bbb7875e596f802296a7a6ce0e77fd72fa51 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 7 Aug 2017 16:28:42 +0200 Subject: Defining abbreviations for tactics that can parse as atoms. --- theories/Notations.v | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/theories/Notations.v b/theories/Notations.v index f11cfbde6e..d7f7170a5e 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -63,6 +63,10 @@ Ltac2 Notation abstract := abstract0. (** Base tactics *) +(** Note that we redeclare notations that can be parsed as mere identifiers + as abbreviations, so that it allows to parse them as function arguments + without having to write them within parentheses. *) + (** Enter and check evar resolution *) Ltac2 enter_h ev f arg := match ev with @@ -76,35 +80,45 @@ Ltac2 intros0 ev p := Control.enter (fun () => Std.intros false p). Ltac2 Notation "intros" p(intropatterns) := intros0 false p. +Ltac2 Notation intros := intros. Ltac2 Notation "eintros" p(intropatterns) := intros0 true p. +Ltac2 Notation eintros := eintros. Ltac2 split0 ev bnd := enter_h ev Std.split bnd. Ltac2 Notation "split" bnd(thunk(bindings)) := split0 false bnd. +Ltac2 Notation split := split. Ltac2 Notation "esplit" bnd(thunk(bindings)) := split0 true bnd. +Ltac2 Notation esplit := esplit. Ltac2 left0 ev bnd := enter_h ev Std.left bnd. Ltac2 Notation "left" bnd(thunk(bindings)) := left0 false bnd. +Ltac2 Notation left := left. Ltac2 Notation "eleft" bnd(thunk(bindings)) := left0 true bnd. +Ltac2 Notation eleft := eleft. Ltac2 right0 ev bnd := enter_h ev Std.right bnd. Ltac2 Notation "right" bnd(thunk(bindings)) := right0 false bnd. +Ltac2 Notation right := right. Ltac2 Notation "eright" bnd(thunk(bindings)) := right0 true bnd. +Ltac2 Notation eright := eright. Ltac2 constructor0 ev n bnd := enter_h ev (fun ev bnd => Std.constructor_n ev n bnd) bnd. Ltac2 Notation "constructor" := Control.enter (fun () => Std.constructor false). +Ltac2 Notation constructor := constructor. Ltac2 Notation "constructor" n(tactic) bnd(thunk(bindings)) := constructor0 false n bnd. Ltac2 Notation "econstructor" := Control.enter (fun () => Std.constructor true). +Ltac2 Notation econstructor := econstructor. Ltac2 Notation "econstructor" n(tactic) bnd(thunk(bindings)) := constructor0 true n bnd. Ltac2 elim0 ev c bnd use := @@ -201,9 +215,11 @@ end. Ltac2 Notation "red" cl(opt(clause)) := Std.red (default_on_concl cl). +Ltac2 Notation red := red. Ltac2 Notation "hnf" cl(opt(clause)) := Std.hnf (default_on_concl cl). +Ltac2 Notation hnf := hnf. Ltac2 rewrite0 ev rw cl tac := let tac := match tac with -- cgit v1.2.3 From 5062231251d58cf51cedb18677392b6e6d168694 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 7 Aug 2017 16:39:25 +0200 Subject: Defining several aliases for built-in tactics. --- tests/example2.v | 4 ++-- theories/Notations.v | 12 ++++++++++++ 2 files changed, 14 insertions(+), 2 deletions(-) diff --git a/tests/example2.v b/tests/example2.v index 76f069a5ae..bfb1b07e7a 100644 --- a/tests/example2.v +++ b/tests/example2.v @@ -5,7 +5,7 @@ Import Ltac2.Notations. Goal exists n, n = 0. Proof. split with (x := 0). -Std.reflexivity (). +reflexivity. Qed. Goal exists n, n = 0. @@ -113,7 +113,7 @@ Qed. Goal forall n m, (m = n -> n = m) -> m = n -> n = 0 -> m = 0. Proof. intros n m He He' He''. -rewrite <- &He by Std.assumption (). +rewrite <- &He by assumption. Control.refine (fun () => &He''). Qed. diff --git a/theories/Notations.v b/theories/Notations.v index d7f7170a5e..0fa3456196 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -242,3 +242,15 @@ Ltac2 Notation "erewrite" cl(opt(clause)) tac(opt(seq("by", thunk(tactic)))) := rewrite0 true rw cl tac. + +(** Other base tactics *) + +Ltac2 Notation reflexivity := Std.reflexivity (). + +Ltac2 Notation assumption := Std.assumption (). + +Ltac2 Notation etransitivity := Std.etransitivity (). + +Ltac2 Notation admit := Std.admit (). + +Ltac2 Notation clear := Std.keep []. -- cgit v1.2.3 From 3d1092cf7df3229ecc2a4da60a33cf7c6b9be1a3 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 8 Aug 2017 12:32:37 +0200 Subject: Code simplification in quotations. --- src/g_ltac2.ml4 | 7 ++++--- src/tac2qexpr.mli | 4 +++- src/tac2quote.ml | 56 +++++++++++++++++++++++++------------------------------ src/tac2quote.mli | 8 ++++---- 4 files changed, 36 insertions(+), 39 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 1cab918ea4..9bc7107cc7 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -325,7 +325,7 @@ GEXTEND Gram [ [ n = Prim.natural -> Loc.tag ~loc:!@loc n ] ] ; q_ident: - [ [ id = ident_or_anti -> Tac2quote.of_anti ~loc:!@loc Tac2quote.of_ident id ] ] + [ [ id = ident_or_anti -> Tac2quote.of_anti Tac2quote.of_ident id ] ] ; qhyp: [ [ x = anti -> x @@ -480,8 +480,9 @@ GEXTEND Gram ] ] ; clause: - [ [ "in"; cl = in_clause -> cl - | "at"; occs = occs_nums -> { q_onhyps = Some []; q_concl_occs = occs } + [ [ "in"; cl = in_clause -> Loc.tag ~loc:!@loc @@ cl + | "at"; occs = occs_nums -> + Loc.tag ~loc:!@loc @@ { q_onhyps = Some []; q_concl_occs = occs } ] ] ; q_clause: diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index f6b8c2c19b..d5520c54ee 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -60,9 +60,11 @@ type occurrences = type hyp_location = (occurrences * Id.t located or_anti) * Locus.hyp_location_flag -type clause = +type clause_r = { q_onhyps : hyp_location list option; q_concl_occs : occurrences; } +type clause = clause_r located + type constr_with_bindings = (Constrexpr.constr_expr * bindings) located type destruction_arg = diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 3e25e1cadb..57c8a4bbed 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -45,9 +45,9 @@ let thunk e = let var = [CPatVar (Some loc, Anonymous), Some (CTypRef (loc, AbsKn (Other t_unit), []))] in CTacFun (loc, var, e) -let of_pair ?loc (e1, e2) = +let of_pair f g (loc, (e1, e2)) = let loc = Option.default dummy_loc loc in - CTacApp (loc, CTacCst (loc, AbsKn (Tuple 2)), [e1; e2]) + CTacApp (loc, CTacCst (loc, AbsKn (Tuple 2)), [f e1; g e2]) let of_tuple ?loc el = let loc = Option.default dummy_loc loc in @@ -57,9 +57,9 @@ let of_tuple ?loc el = let of_int (loc, n) = CTacAtm (Loc.tag ?loc (AtmInt n)) -let of_option ?loc opt = match opt with +let of_option ?loc f opt = match opt with | None -> constructor ?loc (coq_core "None") [] -| Some e -> constructor ?loc (coq_core "Some") [e] +| Some e -> constructor ?loc (coq_core "Some") [f e] let inj_wit ?loc wit x = let loc = Option.default dummy_loc loc in @@ -71,7 +71,7 @@ let of_variable (loc, id) = CErrors.user_err ?loc (str "Invalid identifier") else CTacRef (RelId (Loc.tag ?loc qid)) -let of_anti ?loc f = function +let of_anti f = function | QExpr x -> f x | QAnti id -> of_variable id @@ -89,10 +89,10 @@ let of_bool ?loc b = let c = if b then Core.c_true else Core.c_false in constructor ?loc c [] -let rec of_list ?loc = function +let rec of_list ?loc f = function | [] -> constructor Core.c_nil [] | e :: l -> - constructor ?loc Core.c_cons [e; of_list ?loc l] + constructor ?loc Core.c_cons [f e; of_list ?loc f l] let of_qhyp (loc, h) = match h with | QAnonHyp n -> constructor ?loc Core.c_anon_hyp [of_int n] @@ -102,17 +102,12 @@ let of_bindings (loc, b) = match b with | QNoBindings -> std_constructor ?loc "NoBindings" [] | QImplicitBindings tl -> - let tl = List.map of_open_constr tl in - std_constructor ?loc "ImplicitBindings" [of_list ?loc tl] + std_constructor ?loc "ImplicitBindings" [of_list ?loc of_open_constr tl] | QExplicitBindings tl -> - let map (loc, (qhyp, e)) = of_pair ?loc (of_anti ?loc of_qhyp qhyp, of_open_constr e) in - let tl = List.map map tl in - std_constructor ?loc "ExplicitBindings" [of_list ?loc tl] + let map e = of_pair (fun q -> of_anti of_qhyp q) of_open_constr e in + std_constructor ?loc "ExplicitBindings" [of_list ?loc map tl] -let of_constr_with_bindings (loc, (c, bnd)) = - let c = of_open_constr c in - let bnd = of_bindings bnd in - of_pair ?loc (c, bnd) +let of_constr_with_bindings c = of_pair of_open_constr of_bindings c let rec of_intro_pattern (loc, pat) = match pat with | QIntroForthcoming b -> @@ -126,7 +121,7 @@ and of_intro_pattern_naming (loc, pat) = match pat with | QIntroIdentifier id -> std_constructor ?loc "IntroIdentifier" [of_anti of_ident id] | QIntroFresh id -> - std_constructor ?loc "IntroFresh" [of_anti ?loc of_ident id] + std_constructor ?loc "IntroFresh" [of_anti of_ident id] | QIntroAnonymous -> std_constructor ?loc "IntroAnonymous" [] @@ -142,13 +137,12 @@ and of_intro_pattern_action (loc, pat) = match pat with and of_or_and_intro_pattern (loc, pat) = match pat with | QIntroOrPattern ill -> - let ill = List.map of_intro_patterns ill in - std_constructor ?loc "IntroOrPattern" [of_list ?loc ill] + std_constructor ?loc "IntroOrPattern" [of_list ?loc of_intro_patterns ill] | QIntroAndPattern il -> std_constructor ?loc "IntroAndPattern" [of_intro_patterns il] and of_intro_patterns (loc, l) = - of_list ?loc (List.map of_intro_pattern l) + of_list ?loc of_intro_pattern l let of_hyp_location_flag ?loc = function | Locus.InHyp -> std_constructor ?loc "InHyp" [] @@ -158,25 +152,25 @@ let of_hyp_location_flag ?loc = function let of_occurrences ?loc occ = match occ with | QAllOccurrences -> std_constructor ?loc "AllOccurrences" [] | QAllOccurrencesBut occs -> - let map occ = of_anti ?loc of_int occ in - let occs = of_list ?loc (List.map map occs) in + let map occ = of_anti of_int occ in + let occs = of_list ?loc map occs in std_constructor ?loc "AllOccurrencesBut" [occs] | QNoOccurrences -> std_constructor ?loc "NoOccurrences" [] | QOnlyOccurrences occs -> - let map occ = of_anti ?loc of_int occ in - let occs = of_list ?loc (List.map map occs) in + let map occ = of_anti of_int occ in + let occs = of_list ?loc map occs in std_constructor ?loc "OnlyOccurrences" [occs] let of_hyp_location ?loc ((occs, id), flag) = of_tuple ?loc [ - of_anti ?loc of_ident id; + of_anti of_ident id; of_occurrences ?loc occs; of_hyp_location_flag ?loc flag; ] -let of_clause ?loc cl = +let of_clause (loc, cl) = let loc = Option.default dummy_loc loc in - let hyps = of_option ~loc (Option.map (fun l -> of_list ~loc (List.map of_hyp_location l)) cl.q_onhyps) in + let hyps = of_option ~loc (fun l -> of_list ~loc of_hyp_location l) cl.q_onhyps in let concl = of_occurrences ~loc cl.q_concl_occs in CTacRec (loc, [ std_proj "on_hyps", hyps; @@ -192,9 +186,9 @@ let of_destruction_arg ?loc = function let of_induction_clause (loc, cl) = let arg = of_destruction_arg ?loc cl.indcl_arg in - let eqn = of_option ?loc (Option.map of_intro_pattern_naming cl.indcl_eqn) in - let as_ = of_option ?loc (Option.map of_or_and_intro_pattern cl.indcl_as) in - let in_ = of_option ?loc (Option.map of_clause cl.indcl_in) in + let eqn = of_option ?loc of_intro_pattern_naming cl.indcl_eqn in + let as_ = of_option ?loc of_or_and_intro_pattern cl.indcl_as in + let in_ = of_option ?loc of_clause cl.indcl_in in let loc = Option.default dummy_loc loc in CTacRec (loc, [ std_proj "indcl_arg", arg; @@ -216,7 +210,7 @@ let of_orient loc b = let of_rewriting (loc, rew) = let orient = let (loc, orient) = rew.rew_orient in - of_option ?loc (Option.map (fun b -> of_orient loc b) orient) + of_option ?loc (fun b -> of_orient loc b) orient in let repeat = of_repeat rew.rew_repeat in let equatn = thunk (of_constr_with_bindings rew.rew_equatn) in diff --git a/src/tac2quote.mli b/src/tac2quote.mli index ddb39326d1..dba3c82715 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -19,11 +19,11 @@ open Tac2expr val constructor : ?loc:Loc.t -> ltac_constructor -> raw_tacexpr list -> raw_tacexpr -val of_anti : ?loc:Loc.t -> ('a -> raw_tacexpr) -> 'a or_anti -> raw_tacexpr +val of_anti : ('a -> raw_tacexpr) -> 'a or_anti -> raw_tacexpr val of_int : int located -> raw_tacexpr -val of_pair : ?loc:Loc.t -> raw_tacexpr * raw_tacexpr -> raw_tacexpr +val of_pair : ('a -> raw_tacexpr) -> ('b -> raw_tacexpr) -> ('a * 'b) located -> raw_tacexpr val of_variable : Id.t located -> raw_tacexpr @@ -33,7 +33,7 @@ val of_constr : Constrexpr.constr_expr -> raw_tacexpr val of_open_constr : Constrexpr.constr_expr -> raw_tacexpr -val of_list : ?loc:Loc.t -> raw_tacexpr list -> raw_tacexpr +val of_list : ?loc:Loc.t -> ('a -> raw_tacexpr) -> 'a list -> raw_tacexpr val of_bindings : bindings -> raw_tacexpr @@ -41,7 +41,7 @@ val of_intro_pattern : intro_pattern -> raw_tacexpr val of_intro_patterns : intro_pattern list located -> raw_tacexpr -val of_clause : ?loc:Loc.t -> clause -> raw_tacexpr +val of_clause : clause -> raw_tacexpr val of_induction_clause : induction_clause -> raw_tacexpr -- cgit v1.2.3 From 3fbba861d5355cad92cac52965c8e76a35825c7a Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 8 Aug 2017 14:21:11 +0200 Subject: Another batch of primitive operations. --- src/g_ltac2.ml4 | 8 ++++---- src/tac2core.ml | 2 +- tests/tacticals.v | 22 ++++++++++++++++++++++ theories/Init.v | 3 +++ theories/Notations.v | 28 ++++++++++++++++++++++++++++ 5 files changed, 58 insertions(+), 5 deletions(-) create mode 100644 tests/tacticals.v diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 9bc7107cc7..695fc7d430 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -97,13 +97,13 @@ GEXTEND Gram ] ; tac2expr: - [ "top" RIGHTA + [ "6" RIGHTA [ e1 = SELF; ";"; e2 = SELF -> CTacSeq (!@loc, e1, e2) ] | "5" - [ "fun"; it = LIST1 input_fun ; "=>"; body = tac2expr LEVEL "top" -> CTacFun (!@loc, it, body) + [ "fun"; it = LIST1 input_fun ; "=>"; body = tac2expr LEVEL "6" -> CTacFun (!@loc, it, body) | "let"; isrec = rec_flag; lc = LIST1 let_clause SEP "with"; "in"; - e = tac2expr LEVEL "top" -> CTacLet (!@loc, isrec, lc, e) + e = tac2expr LEVEL "6" -> CTacLet (!@loc, isrec, lc, e) | "match"; e = tac2expr LEVEL "5"; "with"; bl = branches; "end" -> CTacCse (!@loc, e, bl) ] @@ -135,7 +135,7 @@ GEXTEND Gram ] ; branch: - [ [ pat = tac2pat LEVEL "1"; "=>"; e = tac2expr LEVEL "top" -> (pat, e) ] ] + [ [ pat = tac2pat LEVEL "1"; "=>"; e = tac2expr LEVEL "6" -> (pat, e) ] ] ; rec_flag: [ [ IDENT "rec" -> true diff --git a/src/tac2core.ml b/src/tac2core.ml index 08f61f2c6c..1c03cc410d 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -877,7 +877,7 @@ let () = add_scope "tactic" begin function let act tac = tac in Tac2entries.ScopeRule (scope, act) | [SexprInt (loc, n)] -> - let () = if n < 0 || n > 5 then scope_fail () in + let () = if n < 0 || n > 6 then scope_fail () in let scope = Extend.Aentryl (Tac2entries.Pltac.tac2expr, n) in let act tac = tac in Tac2entries.ScopeRule (scope, act) diff --git a/tests/tacticals.v b/tests/tacticals.v new file mode 100644 index 0000000000..73f9c03b87 --- /dev/null +++ b/tests/tacticals.v @@ -0,0 +1,22 @@ +Require Import Ltac2.Ltac2. + +Import Ltac2.Notations. + +Goal True. +Proof. +Fail fail. +Fail solve [ () ]. +try fail. +repeat fail. +repeat (). +solve [ constructor ]. +Qed. + +Goal True. +Proof. +first [ + Message.print (Message.of_string "Yay"); fail +| constructor +| Message.print (Message.of_string "I won't be printed") +]. +Qed. diff --git a/theories/Init.v b/theories/Init.v index 803e48e352..1591747eb4 100644 --- a/theories/Init.v +++ b/theories/Init.v @@ -60,3 +60,6 @@ Ltac2 Type exn ::= [ Not_found ]. Ltac2 Type exn ::= [ Match_failure ]. (** Used to signal a pattern didn't match a term. *) + +Ltac2 Type exn ::= [ Tactic_failure ]. +(** Generic error for tactic failure. *) diff --git a/theories/Notations.v b/theories/Notations.v index 0fa3456196..97f3042d2b 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -27,6 +27,10 @@ match Control.case t with Control.plus (fun _ => s x) (fun e => s (k e)) end. +Ltac2 fail0 (_ : unit) := Control.zero Tactic_failure. + +Ltac2 Notation fail := fail0 (). + Ltac2 try0 t := Control.enter (fun _ => orelse t (fun _ => ())). Ltac2 Notation try := try0. @@ -53,6 +57,28 @@ Ltac2 progress0 tac := Control.enter (fun _ => Control.progress tac). Ltac2 Notation progress := progress0. +Ltac2 rec first0 tacs := +match tacs with +| [] => Control.zero Tactic_failure +| tac :: tacs => Control.enter (fun _ => orelse tac (fun _ => first0 tacs)) +end. + +Ltac2 Notation "first" "[" tacs(list0(thunk(tactic(6)), "|")) "]" := first0 tacs. + +Ltac2 complete tac := + let ans := tac () in + Control.enter (fun () => Control.zero Tactic_failure); + ans. + +Ltac2 rec solve0 tacs := +match tacs with +| [] => Control.zero Tactic_failure +| tac :: tacs => + Control.enter (fun _ => orelse (fun _ => complete tac) (fun _ => first0 tacs)) +end. + +Ltac2 Notation "solve" "[" tacs(list0(thunk(tactic(6)), "|")) "]" := solve0 tacs. + Ltac2 time0 tac := Control.time None tac. Ltac2 Notation time := time0. @@ -254,3 +280,5 @@ Ltac2 Notation etransitivity := Std.etransitivity (). Ltac2 Notation admit := Std.admit (). Ltac2 Notation clear := Std.keep []. + +Ltac2 Notation refine := Control.refine. -- cgit v1.2.3 From 77e3f7be0533fad2c31eb302a51c74b829f99e8c Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 8 Aug 2017 15:54:27 +0200 Subject: Introducing a syntax for goal dispatch. --- src/g_ltac2.ml4 | 20 +++++++++++++++++++- src/tac2core.ml | 1 + src/tac2entries.ml | 1 + src/tac2entries.mli | 1 + src/tac2qexpr.mli | 4 ++++ src/tac2quote.ml | 9 +++++++++ src/tac2quote.mli | 2 ++ tests/tacticals.v | 12 ++++++++++++ theories/Notations.v | 10 ++++++++++ 9 files changed, 59 insertions(+), 1 deletion(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 695fc7d430..e9d2eb1ca3 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -107,6 +107,7 @@ GEXTEND Gram | "match"; e = tac2expr LEVEL "5"; "with"; bl = branches; "end" -> CTacCse (!@loc, e, bl) ] + | "4" LEFTA [ ] | "::" RIGHTA [ e1 = tac2expr; "::"; e2 = tac2expr -> CTacApp (!@loc, CTacCst (!@loc, AbsKn (Other Tac2core.Core.c_cons)), [e1; e2]) @@ -309,7 +310,7 @@ let loc_of_ne_list l = Loc.merge_opt (fst (List.hd l)) (fst (List.last l)) GEXTEND Gram GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause - q_rewriting q_clause; + q_rewriting q_clause q_dispatch; anti: [ [ "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] ; @@ -540,6 +541,23 @@ GEXTEND Gram q_rewriting: [ [ r = oriented_rewriter -> Tac2quote.of_rewriting r ] ] ; + tactic_then_last: + [ [ "|"; lta = LIST0 OPT tac2expr LEVEL "6" SEP "|" -> lta + | -> [] + ] ] + ; + tactic_then_gen: + [ [ ta = tac2expr; "|"; (first,last) = tactic_then_gen -> (Some ta :: first, last) + | ta = tac2expr; ".."; l = tactic_then_last -> ([], Some (Some ta, l)) + | ".."; l = tactic_then_last -> ([], Some (None, l)) + | ta = tac2expr -> ([Some ta], None) + | "|"; (first,last) = tactic_then_gen -> (None :: first, last) + | -> ([None], None) + ] ] + ; + q_dispatch: + [ [ d = tactic_then_gen -> Tac2quote.of_dispatch (Loc.tag ~loc:!@loc d) ] ] + ; END (** Extension of constr syntax *) diff --git a/src/tac2core.ml b/src/tac2core.ml index 1c03cc410d..b8e50ad1df 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -919,6 +919,7 @@ let () = add_expr_scope "intropatterns" Tac2entries.Pltac.q_intropatterns let () = add_expr_scope "induction_clause" Tac2entries.Pltac.q_induction_clause let () = add_expr_scope "rewriting" Tac2entries.Pltac.q_rewriting let () = add_expr_scope "clause" Tac2entries.Pltac.q_clause +let () = add_expr_scope "dispatch" Tac2entries.Pltac.q_dispatch let () = add_generic_scope "constr" Pcoq.Constr.constr Stdarg.wit_constr let () = add_generic_scope "open_constr" Pcoq.Constr.constr Stdarg.wit_open_constr diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 1dd8410d2a..729779aef2 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -31,6 +31,7 @@ let q_intropatterns = Pcoq.Gram.entry_create "tactic:q_intropatterns" let q_induction_clause = Pcoq.Gram.entry_create "tactic:q_induction_clause" let q_rewriting = Pcoq.Gram.entry_create "tactic:q_rewriting" let q_clause = Pcoq.Gram.entry_create "tactic:q_clause" +let q_dispatch = Pcoq.Gram.entry_create "tactic:q_dispatch" end (** Tactic definition *) diff --git a/src/tac2entries.mli b/src/tac2entries.mli index 1fe13cda17..0dd1d5a113 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -64,4 +64,5 @@ val q_intropatterns : raw_tacexpr Pcoq.Gram.entry val q_induction_clause : raw_tacexpr Pcoq.Gram.entry val q_rewriting : raw_tacexpr Pcoq.Gram.entry val q_clause : raw_tacexpr Pcoq.Gram.entry +val q_dispatch : raw_tacexpr Pcoq.Gram.entry end diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index d5520c54ee..e2bf10f4e2 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -96,3 +96,7 @@ type rewriting_r = { } type rewriting = rewriting_r located + +type dispatch_r = raw_tacexpr option list * (raw_tacexpr option * raw_tacexpr option list) option + +type dispatch = dispatch_r located diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 57c8a4bbed..73fd7c29c3 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -230,3 +230,12 @@ let of_exact_hyp ?loc id = let loc = Option.default dummy_loc loc in let refine = CTacRef (AbsKn (TacConstant (control_core "refine"))) in CTacApp (loc, refine, [thunk (of_hyp ~loc id)]) + +let of_dispatch tacs = + let loc = Option.default dummy_loc (fst tacs) in + let default = function + | Some e -> thunk e + | None -> thunk (CTacCst (loc, AbsKn (Tuple 0))) + in + let map e = of_pair default (fun l -> of_list ~loc default l) (Loc.tag ~loc e) in + of_pair (fun l -> of_list ~loc default l) (fun r -> of_option ~loc map r) tacs diff --git a/src/tac2quote.mli b/src/tac2quote.mli index dba3c82715..c02493c554 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -52,3 +52,5 @@ val of_hyp : ?loc:Loc.t -> Id.t located -> raw_tacexpr val of_exact_hyp : ?loc:Loc.t -> Id.t located -> raw_tacexpr (** id ↦ 'Control.refine (fun () => Control.hyp @id') *) + +val of_dispatch : dispatch -> raw_tacexpr diff --git a/tests/tacticals.v b/tests/tacticals.v index 73f9c03b87..1a2fbcbb37 100644 --- a/tests/tacticals.v +++ b/tests/tacticals.v @@ -20,3 +20,15 @@ first [ | Message.print (Message.of_string "I won't be printed") ]. Qed. + +Goal True /\ True. +Proof. +Fail split > [ split | |]. +split > [split | split]. +Qed. + +Goal True /\ (True -> True) /\ True. +Proof. +split > [ | split] > [split | .. | split]. +intros H; refine &H. +Qed. diff --git a/theories/Notations.v b/theories/Notations.v index 97f3042d2b..910b6d5463 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -42,6 +42,16 @@ Ltac2 rec repeat0 (t : unit -> unit) := Ltac2 Notation repeat := repeat0. +Ltac2 dispatch0 t ((head, tail)) := + match tail with + | None => Control.enter (fun _ => t (); Control.dispatch head) + | Some tacs => + let ((def, rem)) := tacs in + Control.enter (fun _ => t (); Control.extend head def rem) + end. + +Ltac2 Notation t(thunk(self)) ">" "[" l(dispatch) "]" : 4 := dispatch0 t l. + Ltac2 do0 n t := let rec aux n t := match Int.equal n 0 with | true => () -- cgit v1.2.3 From 7adc34710bf17c4ec3601831275205c1eb613b84 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 18 Aug 2017 12:47:53 +0200 Subject: Trying to enhance the printing of tactic expressions. --- src/tac2print.ml | 34 +++++++++++++++++----------------- 1 file changed, 17 insertions(+), 17 deletions(-) diff --git a/src/tac2print.ml b/src/tac2print.ml index b679e030fd..e3095c7a89 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -151,14 +151,14 @@ let pr_glbexpr_gen lvl c = | E0 | E1 | E2 | E3 | E4 -> paren | E5 -> fun x -> x in - paren (str "fun" ++ spc () ++ nas ++ spc () ++ str "=>" ++ spc () ++ - hov 0 (pr_glbexpr E5 c)) + paren (hov 0 (hov 2 (str "fun" ++ spc () ++ nas) ++ spc () ++ str "=>" ++ spc () ++ + pr_glbexpr E5 c)) | GTacApp (c, cl) -> let paren = match lvl with | E0 -> paren | E1 | E2 | E3 | E4 | E5 -> fun x -> x in - paren (pr_glbexpr E1 c ++ spc () ++ (pr_sequence (pr_glbexpr E0) cl)) + paren (hov 2 (pr_glbexpr E1 c ++ spc () ++ (pr_sequence (pr_glbexpr E0) cl))) | GTacLet (mut, bnd, e) -> let paren = match lvl with | E0 | E1 | E2 | E3 | E4 -> paren @@ -169,7 +169,7 @@ let pr_glbexpr_gen lvl c = pr_name na ++ spc () ++ str ":=" ++ spc () ++ hov 2 (pr_glbexpr E5 e) ++ spc () in let bnd = prlist_with_sep (fun () -> str "with" ++ spc ()) pr_bnd bnd in - paren (str "let" ++ spc () ++ mut ++ bnd ++ str "in" ++ spc () ++ pr_glbexpr E5 e) + paren (hv 0 (hov 2 (str "let" ++ spc () ++ mut ++ bnd ++ str "in") ++ spc () ++ pr_glbexpr E5 e)) | GTacCst (Tuple 0, _, _) -> str "()" | GTacCst (Tuple _, _, cl) -> let paren = match lvl with @@ -192,7 +192,7 @@ let pr_glbexpr_gen lvl c = | [] -> mt () | _ -> spc () ++ pr_sequence (pr_glbexpr E0) cl in - paren (pr_constructor kn ++ cl) + paren (hov 2 (pr_constructor kn ++ cl)) | _, GTydRec def -> let args = List.combine def cl in let pr_arg ((id, _, _), arg) = @@ -200,7 +200,7 @@ let pr_glbexpr_gen lvl c = pr_projection kn ++ spc () ++ str ":=" ++ spc () ++ pr_glbexpr E1 arg in let args = prlist_with_sep (fun () -> str ";" ++ spc ()) pr_arg args in - str "{" ++ spc () ++ args ++ spc () ++ str "}" + hv 0 (str "{" ++ spc () ++ args ++ spc () ++ str "}") | _, (GTydDef _ | GTydOpn) -> assert false end | GTacCse (e, info, cst_br, ncst_br) -> @@ -219,7 +219,7 @@ let pr_glbexpr_gen lvl c = | [] -> mt () | _ -> spc () ++ pr_sequence pr_name vars in - hov 0 (str "|" ++ spc () ++ cstr ++ vars ++ spc () ++ str "=>" ++ spc () ++ + hov 4 (str "|" ++ spc () ++ hov 0 (cstr ++ vars ++ spc () ++ str "=>") ++ spc () ++ hov 2 (pr_glbexpr E5 p)) ++ spc () in prlist pr_branch br @@ -227,9 +227,9 @@ let pr_glbexpr_gen lvl c = let (vars, p) = if Int.equal n 0 then ([||], cst_br.(0)) else ncst_br.(0) in let p = pr_glbexpr E5 p in let vars = prvect_with_sep (fun () -> str "," ++ spc ()) pr_name vars in - str "|" ++ spc () ++ paren vars ++ spc () ++ str "=>" ++ spc () ++ p + hov 4 (str "|" ++ spc () ++ hov 0 (paren vars ++ spc () ++ str "=>") ++ spc () ++ p) in - hov 0 (hov 0 (str "match" ++ spc () ++ e ++ spc () ++ str "with") ++ spc () ++ Pp.v 0 br ++ str "end") + v 0 (hv 0 (str "match" ++ spc () ++ e ++ spc () ++ str "with") ++ spc () ++ br ++ spc () ++ str "end") | GTacWth wth -> let e = pr_glbexpr E5 wth.opn_match in let pr_pattern c self vars p = @@ -237,7 +237,7 @@ let pr_glbexpr_gen lvl c = | Anonymous -> mt () | Name id -> spc () ++ str "as" ++ spc () ++ Id.print id in - hov 0 (str "|" ++ spc () ++ c ++ vars ++ self ++ spc () ++ str "=>" ++ spc () ++ + hov 4 (str "|" ++ spc () ++ hov 0 (c ++ vars ++ self ++ spc () ++ str "=>") ++ spc () ++ hov 2 (pr_glbexpr E5 p)) ++ spc () in let pr_branch (cstr, (self, vars, p)) = @@ -252,7 +252,7 @@ let pr_glbexpr_gen lvl c = let (def_as, def_p) = wth.opn_default in let def = pr_pattern (str "_") def_as (mt ()) def_p in let br = br ++ def in - hov 0 (hov 0 (str "match" ++ spc () ++ e ++ spc () ++ str "with") ++ spc () ++ Pp.v 0 br ++ str "end") + v 0 (hv 0 (str "match" ++ spc () ++ e ++ spc () ++ str "with") ++ spc () ++ br ++ str "end") | GTacPrj (kn, e, n) -> let def = match Tac2env.interp_type kn with | _, GTydRec def -> def @@ -262,7 +262,7 @@ let pr_glbexpr_gen lvl c = let proj = change_kn_label kn proj in let proj = pr_projection proj in let e = pr_glbexpr E0 e in - e ++ str "." ++ paren proj + hov 0 (e ++ str "." ++ paren proj) | GTacSet (kn, e, n, r) -> let def = match Tac2env.interp_type kn with | _, GTydRec def -> def @@ -273,28 +273,28 @@ let pr_glbexpr_gen lvl c = let proj = pr_projection proj in let e = pr_glbexpr E0 e in let r = pr_glbexpr E1 r in - e ++ str "." ++ paren proj ++ spc () ++ str ":=" ++ spc () ++ r + hov 0 (e ++ str "." ++ paren proj ++ spc () ++ str ":=" ++ spc () ++ r) | GTacOpn (kn, cl) -> let paren = match lvl with | E0 -> paren | E1 | E2 | E3 | E4 | E5 -> fun x -> x in let c = pr_constructor kn in - paren (c ++ spc () ++ (pr_sequence (pr_glbexpr E0) cl)) + paren (hov 0 (c ++ spc () ++ (pr_sequence (pr_glbexpr E0) cl))) | GTacExt arg -> let GenArg (Glbwit tag, arg) = arg in let name = match tag with | ExtraArg tag -> ArgT.repr tag | _ -> assert false in - str name ++ str ":" ++ paren (Genprint.glb_print tag arg) + hov 0 (str name ++ str ":" ++ paren (Genprint.glb_print tag arg)) | GTacPrm (prm, args) -> let args = match args with | [] -> mt () | _ -> spc () ++ pr_sequence (pr_glbexpr E0) args in - str "@external" ++ spc () ++ qstring prm.mltac_plugin ++ spc () ++ - qstring prm.mltac_tactic ++ args + hov 0 (str "@external" ++ spc () ++ qstring prm.mltac_plugin ++ spc () ++ + qstring prm.mltac_tactic ++ args) in hov 0 (pr_glbexpr lvl c) -- cgit v1.2.3 From 900841d0bb4700fb2a3662457e7c4efea34a97e4 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 18 Aug 2017 13:47:36 +0200 Subject: Exporting scopes for occurrences. --- src/g_ltac2.ml4 | 15 +++++++++------ src/tac2core.ml | 2 ++ src/tac2entries.ml | 1 + src/tac2entries.mli | 1 + src/tac2qexpr.mli | 4 +++- src/tac2quote.ml | 6 +++--- src/tac2quote.mli | 2 ++ 7 files changed, 21 insertions(+), 10 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index e9d2eb1ca3..e3d48f75ca 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -310,7 +310,7 @@ let loc_of_ne_list l = Loc.merge_opt (fst (List.hd l)) (fst (List.last l)) GEXTEND Gram GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause - q_rewriting q_clause q_dispatch; + q_rewriting q_clause q_dispatch q_occurrences; anti: [ [ "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] ; @@ -449,13 +449,13 @@ GEXTEND Gram ] ] ; occs_nums: - [ [ nl = LIST1 nat_or_anti -> QOnlyOccurrences nl + [ [ nl = LIST1 nat_or_anti -> Loc.tag ~loc:!@loc @@ QOnlyOccurrences nl | "-"; n = nat_or_anti; nl = LIST0 nat_or_anti -> - QAllOccurrencesBut (n::nl) + Loc.tag ~loc:!@loc @@ QAllOccurrencesBut (n::nl) ] ] ; occs: - [ [ "at"; occs = occs_nums -> occs | -> QAllOccurrences ] ] + [ [ "at"; occs = occs_nums -> occs | -> Loc.tag ~loc:!@loc QAllOccurrences ] ] ; hypident: [ [ id = ident_or_anti -> @@ -477,7 +477,7 @@ GEXTEND Gram | hl = LIST0 hypident_occ SEP ","; "|-"; occs = concl_occ -> { q_onhyps = Some hl; q_concl_occs = occs } | hl = LIST0 hypident_occ SEP "," -> - { q_onhyps = Some hl; q_concl_occs = QNoOccurrences } + { q_onhyps = Some hl; q_concl_occs = Loc.tag ~loc:!@loc QNoOccurrences } ] ] ; clause: @@ -491,7 +491,7 @@ GEXTEND Gram ; concl_occ: [ [ "*"; occs = occs -> occs - | -> QNoOccurrences + | -> Loc.tag ~loc:!@loc QNoOccurrences ] ] ; induction_clause: @@ -558,6 +558,9 @@ GEXTEND Gram q_dispatch: [ [ d = tactic_then_gen -> Tac2quote.of_dispatch (Loc.tag ~loc:!@loc d) ] ] ; + q_occurrences: + [ [ occs = occs -> Tac2quote.of_occurrences occs ] ] + ; END (** Extension of constr syntax *) diff --git a/src/tac2core.ml b/src/tac2core.ml index b8e50ad1df..3ce2ed53a8 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -919,10 +919,12 @@ let () = add_expr_scope "intropatterns" Tac2entries.Pltac.q_intropatterns let () = add_expr_scope "induction_clause" Tac2entries.Pltac.q_induction_clause let () = add_expr_scope "rewriting" Tac2entries.Pltac.q_rewriting let () = add_expr_scope "clause" Tac2entries.Pltac.q_clause +let () = add_expr_scope "occurrences" Tac2entries.Pltac.q_occurrences let () = add_expr_scope "dispatch" Tac2entries.Pltac.q_dispatch let () = add_generic_scope "constr" Pcoq.Constr.constr Stdarg.wit_constr let () = add_generic_scope "open_constr" Pcoq.Constr.constr Stdarg.wit_open_constr +let () = add_generic_scope "pattern" Pcoq.Constr.constr wit_pattern (** seq scope, a bit hairy *) diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 729779aef2..2490f6534b 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -32,6 +32,7 @@ let q_induction_clause = Pcoq.Gram.entry_create "tactic:q_induction_clause" let q_rewriting = Pcoq.Gram.entry_create "tactic:q_rewriting" let q_clause = Pcoq.Gram.entry_create "tactic:q_clause" let q_dispatch = Pcoq.Gram.entry_create "tactic:q_dispatch" +let q_occurrences = Pcoq.Gram.entry_create "tactic:q_occurrences" end (** Tactic definition *) diff --git a/src/tac2entries.mli b/src/tac2entries.mli index 0dd1d5a113..cf6cdfa74b 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -65,4 +65,5 @@ val q_induction_clause : raw_tacexpr Pcoq.Gram.entry val q_rewriting : raw_tacexpr Pcoq.Gram.entry val q_clause : raw_tacexpr Pcoq.Gram.entry val q_dispatch : raw_tacexpr Pcoq.Gram.entry +val q_occurrences : raw_tacexpr Pcoq.Gram.entry end diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index e2bf10f4e2..a631ffd188 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -52,12 +52,14 @@ and intro_pattern_naming = intro_pattern_naming_r located and intro_pattern_action = intro_pattern_action_r located and or_and_intro_pattern = or_and_intro_pattern_r located -type occurrences = +type occurrences_r = | QAllOccurrences | QAllOccurrencesBut of int located or_anti list | QNoOccurrences | QOnlyOccurrences of int located or_anti list +type occurrences = occurrences_r located + type hyp_location = (occurrences * Id.t located or_anti) * Locus.hyp_location_flag type clause_r = diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 73fd7c29c3..b3d174dc2d 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -149,7 +149,7 @@ let of_hyp_location_flag ?loc = function | Locus.InHypTypeOnly -> std_constructor ?loc "InHypTypeOnly" [] | Locus.InHypValueOnly -> std_constructor ?loc "InHypValueOnly" [] -let of_occurrences ?loc occ = match occ with +let of_occurrences (loc, occ) = match occ with | QAllOccurrences -> std_constructor ?loc "AllOccurrences" [] | QAllOccurrencesBut occs -> let map occ = of_anti of_int occ in @@ -164,14 +164,14 @@ let of_occurrences ?loc occ = match occ with let of_hyp_location ?loc ((occs, id), flag) = of_tuple ?loc [ of_anti of_ident id; - of_occurrences ?loc occs; + of_occurrences occs; of_hyp_location_flag ?loc flag; ] let of_clause (loc, cl) = let loc = Option.default dummy_loc loc in let hyps = of_option ~loc (fun l -> of_list ~loc of_hyp_location l) cl.q_onhyps in - let concl = of_occurrences ~loc cl.q_concl_occs in + let concl = of_occurrences cl.q_concl_occs in CTacRec (loc, [ std_proj "on_hyps", hyps; std_proj "on_concl", concl; diff --git a/src/tac2quote.mli b/src/tac2quote.mli index c02493c554..4e563990be 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -47,6 +47,8 @@ val of_induction_clause : induction_clause -> raw_tacexpr val of_rewriting : rewriting -> raw_tacexpr +val of_occurrences : occurrences -> raw_tacexpr + val of_hyp : ?loc:Loc.t -> Id.t located -> raw_tacexpr (** id ↦ 'Control.hyp @id' *) -- cgit v1.2.3 From 62ea702ac88c2762a6587b7b7c95f8f917cedd1c Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 18 Aug 2017 14:40:08 +0200 Subject: Notations for a few reduction functions. --- tests/example2.v | 34 ++++++++++++++++++++++++++++++++++ theories/Notations.v | 18 ++++++++++++++++++ 2 files changed, 52 insertions(+) diff --git a/tests/example2.v b/tests/example2.v index bfb1b07e7a..1856663953 100644 --- a/tests/example2.v +++ b/tests/example2.v @@ -123,3 +123,37 @@ intros n r. hnf in r. split. Qed. + +Goal 1 = 0 -> 0 = 0. +Proof. +intros H. +pattern 0 at 1. +let occ := 2 in pattern 1 at 1, 0 at $occ in H. +reflexivity. +Qed. + +Goal 1 + 1 = 2. +Proof. +vm_compute. +reflexivity. +Qed. + +Goal 1 + 1 = 2. +Proof. +native_compute. +reflexivity. +Qed. + +Goal 1 + 1 = 2 - 0 -> True. +Proof. +intros H. +vm_compute plus in H. +reflexivity. +Qed. + +Goal 1 = 0 -> True /\ True. +Proof. +intros H. +split; fold (1 + 0) (1 + 0) in H. +reflexivity. +Qed. diff --git a/theories/Notations.v b/theories/Notations.v index 910b6d5463..8c2c09a2b5 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -257,6 +257,24 @@ Ltac2 Notation "hnf" cl(opt(clause)) := Std.hnf (default_on_concl cl). Ltac2 Notation hnf := hnf. +Ltac2 Notation "vm_compute" pl(opt(seq(pattern, occurrences))) cl(opt(clause)) := + Std.vm pl (default_on_concl cl). +Ltac2 Notation vm_compute := vm_compute. + +Ltac2 Notation "native_compute" pl(opt(seq(pattern, occurrences))) cl(opt(clause)) := + Std.native pl (default_on_concl cl). +Ltac2 Notation native_compute := native_compute. + +Ltac2 fold0 pl cl := + let cl := default_on_concl cl in + Control.enter (fun () => Control.with_holes pl (fun pl => Std.fold pl cl)). + +Ltac2 Notation "fold" pl(thunk(list1(open_constr))) cl(opt(clause)) := + fold0 pl cl. + +Ltac2 Notation "pattern" pl(list1(seq(constr, occurrences), ",")) cl(opt(clause)) := + Std.pattern pl (default_on_concl cl). + Ltac2 rewrite0 ev rw cl tac := let tac := match tac with | None => None -- cgit v1.2.3 From f392ad50331d3e59d3e2f3ec66c0a42112506d7f Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 18 Aug 2017 16:48:00 +0200 Subject: Laxer dependencies between file and link reordering. --- _CoqProject | 4 ++-- src/ltac2_plugin.mlpack | 2 +- src/tac2core.ml | 4 ---- src/tac2core.mli | 4 ---- src/tac2quote.ml | 15 ++++++--------- 5 files changed, 9 insertions(+), 20 deletions(-) diff --git a/_CoqProject b/_CoqProject index f202e1aed2..4116d17554 100644 --- a/_CoqProject +++ b/_CoqProject @@ -15,11 +15,11 @@ src/tac2entries.ml src/tac2entries.mli src/tac2ffi.ml src/tac2ffi.mli -src/tac2core.ml -src/tac2core.mli src/tac2qexpr.mli src/tac2quote.ml src/tac2quote.mli +src/tac2core.ml +src/tac2core.mli src/tac2tactics.ml src/tac2tactics.mli src/tac2stdlib.ml diff --git a/src/ltac2_plugin.mlpack b/src/ltac2_plugin.mlpack index 4c4082ad65..f9fa2fafd8 100644 --- a/src/ltac2_plugin.mlpack +++ b/src/ltac2_plugin.mlpack @@ -4,8 +4,8 @@ Tac2intern Tac2interp Tac2entries Tac2ffi -Tac2core Tac2quote +Tac2core Tac2tactics Tac2stdlib G_ltac2 diff --git a/src/tac2core.ml b/src/tac2core.ml index 3ce2ed53a8..bf41713215 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -45,10 +45,6 @@ let c_some = coq_core "Some" let c_true = coq_core "true" let c_false = coq_core "false" -let t_qhyp = std_core "hypothesis" -let c_named_hyp = std_core "NamedHyp" -let c_anon_hyp = std_core "AnonHyp" - end open Core diff --git a/src/tac2core.mli b/src/tac2core.mli index 6fd48e85f7..d9ed8ea2e5 100644 --- a/src/tac2core.mli +++ b/src/tac2core.mli @@ -27,8 +27,4 @@ val t_array : type_constant val c_true : ltac_constructor val c_false : ltac_constructor -val t_qhyp : type_constant -val c_anon_hyp : ltac_constructor -val c_named_hyp : ltac_constructor - end diff --git a/src/tac2quote.ml b/src/tac2quote.ml index b3d174dc2d..8adeb15d20 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -9,18 +9,15 @@ open Pp open Names open Util -open Misctypes -open Tac2intern open Tac2expr open Tac2qexpr -open Tac2core (** Syntactic quoting of expressions. *) let control_prefix = MPfile (DirPath.make (List.map Id.of_string ["Control"; "Ltac2"])) -let kername prefix n = KerName.make2 prefix (Label.of_id (Id.of_string n)) +let kername prefix n = KerName.make2 prefix (Label.of_id (Id.of_string_soft n)) let std_core n = kername Tac2env.std_prefix n let coq_core n = kername Tac2env.coq_prefix n let control_core n = kername control_prefix n @@ -86,17 +83,17 @@ let of_open_constr c = inj_wit ?loc Stdarg.wit_open_constr c let of_bool ?loc b = - let c = if b then Core.c_true else Core.c_false in + let c = if b then coq_core "true" else coq_core "false" in constructor ?loc c [] let rec of_list ?loc f = function -| [] -> constructor Core.c_nil [] +| [] -> constructor (coq_core "[]") [] | e :: l -> - constructor ?loc Core.c_cons [f e; of_list ?loc f l] + constructor ?loc (coq_core "::") [f e; of_list ?loc f l] let of_qhyp (loc, h) = match h with -| QAnonHyp n -> constructor ?loc Core.c_anon_hyp [of_int n] -| QNamedHyp id -> constructor ?loc Core.c_named_hyp [of_ident id] +| QAnonHyp n -> std_constructor ?loc "AnonHyp" [of_int n] +| QNamedHyp id -> std_constructor ?loc "NamedHyp" [of_ident id] let of_bindings (loc, b) = match b with | QNoBindings -> -- cgit v1.2.3 From 33f7df93bb686077b9ca164078763c2208cbe3d5 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 18 Aug 2017 17:03:37 +0200 Subject: Removing dead code. --- src/tac2core.ml | 1 - src/tac2core.mli | 2 -- src/tac2intern.ml | 1 - src/tac2intern.mli | 1 - src/tac2interp.mli | 1 - src/tac2print.mli | 2 -- src/tac2qexpr.mli | 1 - src/tac2quote.mli | 1 - src/tac2stdlib.ml | 8 -------- src/tac2tactics.ml | 4 ---- 10 files changed, 22 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index bf41713215..0415b6f15f 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -21,7 +21,6 @@ open Proofview.Notations module Value = Tac2ffi let coq_core n = KerName.make2 Tac2env.coq_prefix (Label.of_id (Id.of_string_soft n)) -let std_core n = KerName.make2 Tac2env.std_prefix (Label.of_id (Id.of_string_soft n)) module Core = struct diff --git a/src/tac2core.mli b/src/tac2core.mli index d9ed8ea2e5..566b7efbb7 100644 --- a/src/tac2core.mli +++ b/src/tac2core.mli @@ -6,8 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Names -open Tac2env open Tac2expr (** {5 Hardwired data} *) diff --git a/src/tac2intern.ml b/src/tac2intern.ml index bf7e93cb9e..b62a574a6c 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -24,7 +24,6 @@ let coq_type n = KerName.make2 Tac2env.coq_prefix (Label.make n) let t_int = coq_type "int" let t_string = coq_type "string" let t_array = coq_type "array" -let t_unit = coq_type "unit" let t_list = coq_type "list" let c_nil = GTacCst (Other t_list, 0, []) diff --git a/src/tac2intern.mli b/src/tac2intern.mli index ddec8eb7e4..898df649ba 100644 --- a/src/tac2intern.mli +++ b/src/tac2intern.mli @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Genarg open Names open Mod_subst open Tac2expr diff --git a/src/tac2interp.mli b/src/tac2interp.mli index bf6b2d4dde..42e9e3adeb 100644 --- a/src/tac2interp.mli +++ b/src/tac2interp.mli @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Genarg open Names open Tac2expr diff --git a/src/tac2print.mli b/src/tac2print.mli index 2ee5cf42e0..0a04af2ff0 100644 --- a/src/tac2print.mli +++ b/src/tac2print.mli @@ -6,8 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Pp -open Names open Tac2expr (** {5 Printing types} *) diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index a631ffd188..2e7fc7b44a 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Util open Loc open Names open Tac2expr diff --git a/src/tac2quote.mli b/src/tac2quote.mli index 4e563990be..456d1fa97d 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -8,7 +8,6 @@ open Loc open Names -open Misctypes open Tac2qexpr open Tac2expr diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 3cfd0b5626..e3b0d6bf6b 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -11,8 +11,6 @@ open Locus open Misctypes open Genredexpr open Tac2expr -open Tac2core -open Tac2tactics open Proofview.Notations module Value = Tac2ffi @@ -169,12 +167,6 @@ let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } let lift tac = tac <*> return v_unit -let wrap f = - return () >>= fun () -> return (f ()) - -let wrap_unit f = - return () >>= fun () -> f (); return v_unit - let define_prim0 name tac = let tac = function | [_] -> lift tac diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index 50c1df922e..7fdda1f17d 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -12,11 +12,7 @@ open Globnames open Misctypes open Tactypes open Genredexpr -open Tactics open Proofview -open Tacmach.New -open Tacticals.New -open Proofview.Notations (** FIXME: export a better interface in Tactics *) let delayed_of_tactic tac env sigma = -- cgit v1.2.3 From 0b2c0e58b45b2e78f8ad65ddbc7254e1fd9d07eb Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 18 Aug 2017 17:07:14 +0200 Subject: More precise type for quotation entries. --- src/g_ltac2.ml4 | 18 +++++++++--------- src/tac2core.ml | 41 ++++++++++++++--------------------------- src/tac2entries.mli | 24 +++++++++++++----------- 3 files changed, 36 insertions(+), 47 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index e3d48f75ca..e5847119e1 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -326,7 +326,7 @@ GEXTEND Gram [ [ n = Prim.natural -> Loc.tag ~loc:!@loc n ] ] ; q_ident: - [ [ id = ident_or_anti -> Tac2quote.of_anti Tac2quote.of_ident id ] ] + [ [ id = ident_or_anti -> id ] ] ; qhyp: [ [ x = anti -> x @@ -347,7 +347,7 @@ GEXTEND Gram ] ] ; q_bindings: - [ [ bl = with_bindings -> Tac2quote.of_bindings bl ] ] + [ [ bl = with_bindings -> bl ] ] ; intropatterns: [ [ l = LIST0 nonsimple_intropattern -> Loc.tag ~loc:!@loc l ]] @@ -417,10 +417,10 @@ GEXTEND Gram ] ] ; q_intropatterns: - [ [ ipat = intropatterns -> Tac2quote.of_intro_patterns ipat ] ] + [ [ ipat = intropatterns -> ipat ] ] ; q_intropattern: - [ [ ipat = simple_intropattern -> Tac2quote.of_intro_pattern ipat ] ] + [ [ ipat = simple_intropattern -> ipat ] ] ; nat_or_anti: [ [ n = lnatural -> QExpr n @@ -487,7 +487,7 @@ GEXTEND Gram ] ] ; q_clause: - [ [ cl = clause -> Tac2quote.of_clause cl ] ] + [ [ cl = clause -> cl ] ] ; concl_occ: [ [ "*"; occs = occs -> occs @@ -506,7 +506,7 @@ GEXTEND Gram ] ] ; q_induction_clause: - [ [ cl = induction_clause -> Tac2quote.of_induction_clause cl ] ] + [ [ cl = induction_clause -> cl ] ] ; orient: [ [ "->" -> Loc.tag ~loc:!@loc (Some true) @@ -539,7 +539,7 @@ GEXTEND Gram ] ] ; q_rewriting: - [ [ r = oriented_rewriter -> Tac2quote.of_rewriting r ] ] + [ [ r = oriented_rewriter -> r ] ] ; tactic_then_last: [ [ "|"; lta = LIST0 OPT tac2expr LEVEL "6" SEP "|" -> lta @@ -556,10 +556,10 @@ GEXTEND Gram ] ] ; q_dispatch: - [ [ d = tactic_then_gen -> Tac2quote.of_dispatch (Loc.tag ~loc:!@loc d) ] ] + [ [ d = tactic_then_gen -> Loc.tag ~loc:!@loc d ] ] ; q_occurrences: - [ [ occs = occs -> Tac2quote.of_occurrences occs ] ] + [ [ occs = occs -> occs ] ] ; END diff --git a/src/tac2core.ml b/src/tac2core.ml index 0415b6f15f..bec1761120 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -14,6 +14,7 @@ open Geninterp open Tac2env open Tac2expr open Tac2interp +open Tac2entries.Pltac open Proofview.Notations (** Standard values *) @@ -868,20 +869,12 @@ end let () = add_scope "tactic" begin function | [] -> (** Default to level 5 parsing *) - let scope = Extend.Aentryl (Tac2entries.Pltac.tac2expr, 5) in + let scope = Extend.Aentryl (tac2expr, 5) in let act tac = tac in Tac2entries.ScopeRule (scope, act) | [SexprInt (loc, n)] -> let () = if n < 0 || n > 6 then scope_fail () in - let scope = Extend.Aentryl (Tac2entries.Pltac.tac2expr, n) in - let act tac = tac in - Tac2entries.ScopeRule (scope, act) -| _ -> scope_fail () -end - -let () = add_scope "ident" begin function -| [] -> - let scope = Extend.Aentry Tac2entries.Pltac.q_ident in + let scope = Extend.Aentryl (tac2expr, n) in let act tac = tac in Tac2entries.ScopeRule (scope, act) | _ -> scope_fail () @@ -895,27 +888,21 @@ let () = add_scope "thunk" begin function | _ -> scope_fail () end -let () = add_scope "bindings" begin function -| [] -> - let scope = Extend.Aentry Tac2entries.Pltac.q_bindings in - let act tac = tac in - Tac2entries.ScopeRule (scope, act) -| _ -> scope_fail () -end - -let add_expr_scope name entry = +let add_expr_scope name entry f = add_scope name begin function - | [] -> Tac2entries.ScopeRule (Extend.Aentry entry, (fun e -> e)) + | [] -> Tac2entries.ScopeRule (Extend.Aentry entry, f) | _ -> scope_fail () end -let () = add_expr_scope "intropattern" Tac2entries.Pltac.q_intropattern -let () = add_expr_scope "intropatterns" Tac2entries.Pltac.q_intropatterns -let () = add_expr_scope "induction_clause" Tac2entries.Pltac.q_induction_clause -let () = add_expr_scope "rewriting" Tac2entries.Pltac.q_rewriting -let () = add_expr_scope "clause" Tac2entries.Pltac.q_clause -let () = add_expr_scope "occurrences" Tac2entries.Pltac.q_occurrences -let () = add_expr_scope "dispatch" Tac2entries.Pltac.q_dispatch +let () = add_expr_scope "ident" q_ident (fun id -> Tac2quote.of_anti Tac2quote.of_ident id) +let () = add_expr_scope "bindings" q_bindings Tac2quote.of_bindings +let () = add_expr_scope "intropattern" q_intropattern Tac2quote.of_intro_pattern +let () = add_expr_scope "intropatterns" q_intropatterns Tac2quote.of_intro_patterns +let () = add_expr_scope "induction_clause" q_induction_clause Tac2quote.of_induction_clause +let () = add_expr_scope "rewriting" q_rewriting Tac2quote.of_rewriting +let () = add_expr_scope "clause" q_clause Tac2quote.of_clause +let () = add_expr_scope "occurrences" q_occurrences Tac2quote.of_occurrences +let () = add_expr_scope "dispatch" q_dispatch Tac2quote.of_dispatch let () = add_generic_scope "constr" Pcoq.Constr.constr Stdarg.wit_constr let () = add_generic_scope "open_constr" Pcoq.Constr.constr Stdarg.wit_open_constr diff --git a/src/tac2entries.mli b/src/tac2entries.mli index cf6cdfa74b..667378030a 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -55,15 +55,17 @@ module Pltac : sig val tac2expr : raw_tacexpr Pcoq.Gram.entry -(** Quoted entries. They directly return an Ltac2 expression *) - -val q_ident : raw_tacexpr Pcoq.Gram.entry -val q_bindings : raw_tacexpr Pcoq.Gram.entry -val q_intropattern : raw_tacexpr Pcoq.Gram.entry -val q_intropatterns : raw_tacexpr Pcoq.Gram.entry -val q_induction_clause : raw_tacexpr Pcoq.Gram.entry -val q_rewriting : raw_tacexpr Pcoq.Gram.entry -val q_clause : raw_tacexpr Pcoq.Gram.entry -val q_dispatch : raw_tacexpr Pcoq.Gram.entry -val q_occurrences : raw_tacexpr Pcoq.Gram.entry +(** Quoted entries. To be used for complex notations. *) + +open Tac2qexpr + +val q_ident : Id.t located or_anti Pcoq.Gram.entry +val q_bindings : bindings Pcoq.Gram.entry +val q_intropattern : intro_pattern Pcoq.Gram.entry +val q_intropatterns : intro_pattern list located Pcoq.Gram.entry +val q_induction_clause : induction_clause Pcoq.Gram.entry +val q_rewriting : rewriting Pcoq.Gram.entry +val q_clause : clause Pcoq.Gram.entry +val q_dispatch : dispatch Pcoq.Gram.entry +val q_occurrences : occurrences Pcoq.Gram.entry end -- cgit v1.2.3 From 0232b0de849998d3394a4e6a2ab6232a75897610 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 18 Aug 2017 17:59:49 +0200 Subject: Use references in reduction tactics. We dynamically check that the provided references are indeed evaluable ones, instead of ensuring this at internalization time. --- src/tac2stdlib.ml | 24 ++++++++++++------------ src/tac2tactics.ml | 34 ++++++++++++++++++++++++++++++++++ src/tac2tactics.mli | 11 ++++++++++- theories/Std.v | 12 +++++++----- 4 files changed, 63 insertions(+), 18 deletions(-) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index e3b0d6bf6b..d3430213b4 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -8,6 +8,7 @@ open Names open Locus +open Globnames open Misctypes open Genredexpr open Tac2expr @@ -70,9 +71,11 @@ let to_clause = function { onhyps = hyps; concl_occs = to_occurrences to_int_or_var concl; } | _ -> assert false -let to_evaluable_ref = function -| ValBlk (0, [| id |]) -> EvalVarRef (Value.to_ident id) -| ValBlk (1, [| cst |]) -> EvalConstRef (Value.to_constant cst) +let to_reference = function +| ValBlk (0, [| id |]) -> VarRef (Value.to_ident id) +| ValBlk (1, [| cst |]) -> ConstRef (Value.to_constant cst) +| ValBlk (2, [| ind |]) -> IndRef (Value.to_ext Value.val_inductive ind) +| ValBlk (3, [| cstr |]) -> ConstructRef (Value.to_ext Value.val_constructor cstr) | _ -> assert false let to_red_flag = function @@ -84,7 +87,7 @@ let to_red_flag = function rCofix = Value.to_bool cofix; rZeta = Value.to_bool zeta; rDelta = Value.to_bool delta; - rConst = Value.to_list to_evaluable_ref const; + rConst = Value.to_list to_reference const; } | _ -> assert false @@ -310,29 +313,26 @@ end let () = define_prim2 "tac_cbv" begin fun flags cl -> let flags = to_red_flag flags in let cl = to_clause cl in - Tactics.reduce (Cbv flags) cl + Tac2tactics.cbv flags cl end let () = define_prim2 "tac_cbn" begin fun flags cl -> let flags = to_red_flag flags in let cl = to_clause cl in - Tactics.reduce (Cbn flags) cl + Tac2tactics.cbn flags cl end let () = define_prim2 "tac_lazy" begin fun flags cl -> let flags = to_red_flag flags in let cl = to_clause cl in - Tactics.reduce (Lazy flags) cl + Tac2tactics.lazy_ flags cl end let () = define_prim2 "tac_unfold" begin fun refs cl -> - let map v = - let (ref, occ) = to_pair to_evaluable_ref (fun occ -> to_occurrences to_int_or_var occ) v in - (occ, ref) - in + let map v = to_pair to_reference (fun occ -> to_occurrences to_int_or_var occ) v in let refs = Value.to_list map refs in let cl = to_clause cl in - Tactics.reduce (Unfold refs) cl + Tac2tactics.unfold refs cl end let () = define_prim2 "tac_fold" begin fun args cl -> diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index 7fdda1f17d..2f4965783f 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open Pp open Util open Names open Globnames @@ -13,6 +14,7 @@ open Misctypes open Tactypes open Genredexpr open Proofview +open Proofview.Notations (** FIXME: export a better interface in Tactics *) let delayed_of_tactic tac env sigma = @@ -67,10 +69,42 @@ let map_pattern_with_occs (pat, occ) = match pat with | Pattern.PRef (VarRef id) -> (occ, Inl (EvalVarRef id)) | _ -> (occ, Inr pat) +let get_evaluable_reference = function +| VarRef id -> Proofview.tclUNIT (EvalVarRef id) +| ConstRef cst -> Proofview.tclUNIT (EvalConstRef cst) +| r -> + Tacticals.New.tclZEROMSG (str "Cannot coerce" ++ spc () ++ + Nametab.pr_global_env Id.Set.empty r ++ spc () ++ + str "to an evaluable reference.") + let simpl flags where cl = let where = Option.map map_pattern_with_occs where in + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in Tactics.reduce (Simpl (flags, where)) cl +let cbv flags cl = + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + Tactics.reduce (Cbv flags) cl + +let cbn flags cl = + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + Tactics.reduce (Cbn flags) cl + +let lazy_ flags cl = + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + Tactics.reduce (Lazy flags) cl + +let unfold occs cl = + let map (gr, occ) = + get_evaluable_reference gr >>= fun gr -> Proofview.tclUNIT (occ, gr) + in + Proofview.Monad.List.map map occs >>= fun occs -> + Tactics.reduce (Unfold occs) cl + let vm where cl = let where = Option.map map_pattern_with_occs where in Tactics.reduce (CbvVm where) cl diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli index affbbbbdd7..d835f768a1 100644 --- a/src/tac2tactics.mli +++ b/src/tac2tactics.mli @@ -8,6 +8,7 @@ open Names open Locus +open Globnames open Genredexpr open Misctypes open Tactypes @@ -36,9 +37,17 @@ type rewriting = val rewrite : evars_flag -> rewriting list -> clause -> unit tactic option -> unit tactic -val simpl : evaluable_global_reference glob_red_flag -> +val simpl : global_reference glob_red_flag -> (Pattern.constr_pattern * occurrences_expr) option -> clause -> unit tactic +val cbv : global_reference glob_red_flag -> clause -> unit tactic + +val cbn : global_reference glob_red_flag -> clause -> unit tactic + +val lazy_ : global_reference glob_red_flag -> clause -> unit tactic + +val unfold : (global_reference * occurrences_expr) list -> clause -> unit tactic + val vm : (Pattern.constr_pattern * occurrences_expr) option -> clause -> unit tactic val native : (Pattern.constr_pattern * occurrences_expr) option -> clause -> unit tactic diff --git a/theories/Std.v b/theories/Std.v index 43ccb06192..dd81835c40 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -34,9 +34,11 @@ Ltac2 Type clause := { on_concl : occurrences; }. -Ltac2 Type evaluable_reference := [ -| EvalVarRef (ident) -| EvalConstRef (constant) +Ltac2 Type reference := [ +| VarRef (ident) +| ConstRef (constant) +| IndRef (inductive) +| ConstructRef (constructor) ]. Ltac2 Type red_flags := { @@ -46,7 +48,7 @@ Ltac2 Type red_flags := { rCofix : bool; rZeta : bool; rDelta : bool; (** true = delta all but rConst; false = delta only on rConst*) - rConst : evaluable_reference list + rConst : reference list }. Ltac2 Type 'a not_implemented. @@ -134,7 +136,7 @@ Ltac2 @ external simpl : red_flags -> (pattern * occurrences) option -> clause - Ltac2 @ external cbv : red_flags -> clause -> unit := "ltac2" "tac_cbv". Ltac2 @ external cbn : red_flags -> clause -> unit := "ltac2" "tac_cbn". Ltac2 @ external lazy : red_flags -> clause -> unit := "ltac2" "tac_lazy". -Ltac2 @ external unfold : (evaluable_reference * occurrences) list -> clause -> unit := "ltac2" "tac_unfold". +Ltac2 @ external unfold : (reference * occurrences) list -> clause -> unit := "ltac2" "tac_unfold". Ltac2 @ external fold : constr list -> clause -> unit := "ltac2" "tac_fold". Ltac2 @ external pattern : (constr * occurrences) list -> clause -> unit := "ltac2" "tac_pattern". Ltac2 @ external vm : (pattern * occurrences) option -> clause -> unit := "ltac2" "tac_vm". -- cgit v1.2.3 From 7d496e618f35a17b8432ac3c7205138f03c95fd2 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 24 Aug 2017 14:58:46 +0200 Subject: Introducing a quotation for global references. --- doc/ltac2.md | 4 ++++ src/g_ltac2.ml4 | 7 +++++++ src/tac2core.ml | 33 ++++++++++++++++++++++++++++++--- src/tac2env.ml | 5 +++-- src/tac2env.mli | 9 +++++++-- src/tac2ffi.ml | 14 ++++++++++++++ src/tac2ffi.mli | 3 +++ src/tac2interp.ml | 2 +- src/tac2stdlib.ml | 11 ++--------- tests/quot.v | 9 +++++++++ 10 files changed, 80 insertions(+), 17 deletions(-) create mode 100644 tests/quot.v diff --git a/doc/ltac2.md b/doc/ltac2.md index b3596b2977..6c4912c8f3 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -437,6 +437,10 @@ The current implementation recognizes the following built-in quotations: holes at runtime (type `Init.constr` as well). - "pattern", which parses Coq patterns and produces a pattern used for term matching (type `Init.pattern`). +- "reference", which parses either a `QUALID` or `"&" IDENT`. Qualified names + are globalized at internalization into the corresponding global reference, + while `&id` is turned into `Std.VarRef id`. This produces at runtime a + `Std.reference`. The following syntactic sugar is provided for two common cases. - `@id` is the same as ident:(id) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index e5847119e1..c70f1e882d 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -65,6 +65,7 @@ let tac2mode = Gram.entry_create "vernac:ltac2_command" let inj_wit wit loc x = CTacExt (loc, Genarg.in_gen (Genarg.rawwit wit) x) let inj_open_constr loc c = inj_wit Stdarg.wit_open_constr loc c let inj_pattern loc c = inj_wit Tac2env.wit_pattern loc c +let inj_reference loc c = inj_wit Tac2env.wit_reference loc c let pattern_of_qualid loc id = if Tac2env.is_constructor (snd id) then CPatRef (loc, RelId id, []) @@ -157,6 +158,7 @@ GEXTEND Gram | IDENT "open_constr"; ":"; "("; c = Constr.lconstr; ")" -> Tac2quote.of_open_constr c | IDENT "ident"; ":"; "("; c = lident; ")" -> Tac2quote.of_ident c | IDENT "pattern"; ":"; "("; c = Constr.lconstr_pattern; ")" -> inj_pattern !@loc c + | IDENT "reference"; ":"; "("; c = globref; ")" -> inj_reference !@loc c ] ] ; let_clause: @@ -300,6 +302,11 @@ GEXTEND Gram lident: [ [ id = Prim.ident -> Loc.tag ~loc:!@loc id ] ] ; + globref: + [ [ "&"; id = Prim.ident -> Libnames.Ident (Loc.tag ~loc:!@loc id) + | qid = Prim.qualid -> Libnames.Qualid qid + ] ] + ; END (** Quotation scopes used by notations *) diff --git a/src/tac2core.ml b/src/tac2core.ml index bec1761120..303d62a30d 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -689,7 +689,7 @@ let interp_constr flags ist (c, _) = Proofview.V82.wrap_exceptions begin fun () -> let ist = to_lvar ist in let (sigma, c) = understand_ltac flags env sigma ist WithoutTypeConstraint c in - let c = Val.Dyn (Value.val_constr, c) in + let c = ValExt (Val.Dyn (Value.val_constr, c)) in Proofview.Unsafe.tclEVARS sigma >>= fun () -> Proofview.tclUNIT c end @@ -712,7 +712,7 @@ let () = define_ml_object Stdarg.wit_open_constr obj let () = - let interp _ id = return (Val.Dyn (Value.val_ident, id)) in + let interp _ id = return (ValExt (Val.Dyn (Value.val_ident, id))) in let obj = { ml_type = t_ident; ml_interp = interp; @@ -720,13 +720,21 @@ let () = define_ml_object Stdarg.wit_ident obj let () = - let interp _ c = return (Val.Dyn (Value.val_pattern, c)) in + let interp _ c = return (ValExt (Val.Dyn (Value.val_pattern, c))) in let obj = { ml_type = t_pattern; ml_interp = interp; } in define_ml_object Tac2env.wit_pattern obj +let () = + let interp _ gr = return (Value.of_reference gr) in + let obj = { + ml_type = t_pattern; + ml_interp = interp; + } in + define_ml_object Tac2env.wit_reference obj + let () = let interp ist env sigma concl tac = let fold id (Val.Dyn (tag, v)) (accu : environment) : environment = @@ -754,6 +762,25 @@ let () = let subst s c = Patternops.subst_pattern s c in Genintern.register_subst0 wit_pattern subst +(** References *) + +let () = + let intern ist qid = match qid with + | Libnames.Ident (_, id) -> ist, Globnames.VarRef id + | Libnames.Qualid (loc, qid) -> + let gr = + try Nametab.locate qid + with Not_found -> + Nametab.error_global_not_found ?loc qid + in + ist, gr + in + Genintern.register_intern0 wit_reference intern + +let () = + let subst s c = Globnames.subst_global_reference s c in + Genintern.register_subst0 wit_reference subst + (** Built-in notation scopes *) let add_scope s f = diff --git a/src/tac2env.ml b/src/tac2env.ml index ac2bd5fc23..65276ec57f 100644 --- a/src/tac2env.ml +++ b/src/tac2env.ml @@ -246,7 +246,7 @@ let shortest_qualid_of_projection kn = type 'a ml_object = { ml_type : type_constant; - ml_interp : environment -> 'a -> Geninterp.Val.t Proofview.tactic; + ml_interp : environment -> 'a -> valexpr Proofview.tactic; } module MLTypeObj = @@ -271,8 +271,9 @@ let std_prefix = (** Generic arguments *) -let wit_ltac2 = Genarg.make0 "ltac2" +let wit_ltac2 = Genarg.make0 "ltac2:value" let wit_pattern = Genarg.make0 "ltac2:pattern" +let wit_reference = Genarg.make0 "ltac2:reference" let is_constructor qid = let (_, id) = repr_qualid qid in diff --git a/src/tac2env.mli b/src/tac2env.mli index e4cc8387c5..20bf24d19d 100644 --- a/src/tac2env.mli +++ b/src/tac2env.mli @@ -100,7 +100,7 @@ val interp_primitive : ml_tactic_name -> ml_tactic type 'a ml_object = { ml_type : type_constant; - ml_interp : environment -> 'a -> Geninterp.Val.t Proofview.tactic; + ml_interp : environment -> 'a -> valexpr Proofview.tactic; } val define_ml_object : ('a, 'b, 'c) genarg_type -> 'b ml_object -> unit @@ -118,7 +118,12 @@ val std_prefix : ModPath.t val wit_ltac2 : (raw_tacexpr, glb_tacexpr, Util.Empty.t) genarg_type -val wit_pattern : (Constrexpr.constr_expr, Pattern.constr_pattern, Pattern.constr_pattern) genarg_type +val wit_pattern : (Constrexpr.constr_expr, Pattern.constr_pattern, Util.Empty.t) genarg_type + +val wit_reference : (reference, Globnames.global_reference, Util.Empty.t) genarg_type +(** Beware, at the raw level, [Qualid [id]] has not the same meaning as + [Ident id]. The first is an unqualified global reference, the second is + the dynamic reference to id. *) (** {5 Helper functions} *) diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index 49b49d92fd..b506a578b1 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -7,6 +7,7 @@ (************************************************************************) open Util +open Globnames open Genarg open Geninterp open Tac2expr @@ -125,3 +126,16 @@ let to_array f = function let of_constant c = of_ext val_constant c let to_constant c = to_ext val_constant c + +let of_reference = function +| VarRef id -> ValBlk (0, [| of_ident id |]) +| ConstRef cst -> ValBlk (1, [| of_constant cst |]) +| IndRef ind -> ValBlk (2, [| of_ext val_inductive ind |]) +| ConstructRef cstr -> ValBlk (3, [| of_ext val_constructor cstr |]) + +let to_reference = function +| ValBlk (0, [| id |]) -> VarRef (to_ident id) +| ValBlk (1, [| cst |]) -> ConstRef (to_constant cst) +| ValBlk (2, [| ind |]) -> IndRef (to_ext val_inductive ind) +| ValBlk (3, [| cstr |]) -> ConstructRef (to_ext val_constructor cstr) +| _ -> assert false diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index b69ca9a382..71d90ba940 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -62,6 +62,9 @@ val to_pp : valexpr -> Pp.t val of_constant : Constant.t -> valexpr val to_constant : valexpr -> Constant.t +val of_reference : Globnames.global_reference -> valexpr +val to_reference : valexpr -> Globnames.global_reference + val of_ext : 'a Val.typ -> 'a -> valexpr val to_ext : 'a Val.typ -> valexpr -> 'a diff --git a/src/tac2interp.ml b/src/tac2interp.ml index 664b7de3d6..d3bc79957b 100644 --- a/src/tac2interp.ml +++ b/src/tac2interp.ml @@ -105,7 +105,7 @@ let rec interp ist = function | GTacExt e -> let GenArg (Glbwit tag, e) = e in let tpe = Tac2env.interp_ml_object tag in - tpe.Tac2env.ml_interp ist e >>= fun e -> return (ValExt e) + tpe.Tac2env.ml_interp ist e and interp_app f args = match f with | ValCls { clos_env = ist; clos_var = ids; clos_exp = e } -> diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index d3430213b4..eccaf95ab3 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -71,13 +71,6 @@ let to_clause = function { onhyps = hyps; concl_occs = to_occurrences to_int_or_var concl; } | _ -> assert false -let to_reference = function -| ValBlk (0, [| id |]) -> VarRef (Value.to_ident id) -| ValBlk (1, [| cst |]) -> ConstRef (Value.to_constant cst) -| ValBlk (2, [| ind |]) -> IndRef (Value.to_ext Value.val_inductive ind) -| ValBlk (3, [| cstr |]) -> ConstructRef (Value.to_ext Value.val_constructor cstr) -| _ -> assert false - let to_red_flag = function | ValBlk (0, [| beta; iota; fix; cofix; zeta; delta; const |]) -> { @@ -87,7 +80,7 @@ let to_red_flag = function rCofix = Value.to_bool cofix; rZeta = Value.to_bool zeta; rDelta = Value.to_bool delta; - rConst = Value.to_list to_reference const; + rConst = Value.to_list Value.to_reference const; } | _ -> assert false @@ -329,7 +322,7 @@ let () = define_prim2 "tac_lazy" begin fun flags cl -> end let () = define_prim2 "tac_unfold" begin fun refs cl -> - let map v = to_pair to_reference (fun occ -> to_occurrences to_int_or_var occ) v in + let map v = to_pair Value.to_reference (fun occ -> to_occurrences to_int_or_var occ) v in let refs = Value.to_list map refs in let cl = to_clause cl in Tac2tactics.unfold refs cl diff --git a/tests/quot.v b/tests/quot.v new file mode 100644 index 0000000000..c9aa1f9d14 --- /dev/null +++ b/tests/quot.v @@ -0,0 +1,9 @@ +Require Import Ltac2.Ltac2. + +(** Test for quotations *) + +Ltac2 ref0 () := reference:(&x). +Ltac2 ref1 () := reference:(nat). +Ltac2 ref2 () := reference:(Datatypes.nat). +Fail Ltac2 ref () := reference:(i_certainly_dont_exist). +Fail Ltac2 ref () := reference:(And.Me.neither). -- cgit v1.2.3 From d61047ba240741b9f92ec03e7026deba79ea7b70 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 24 Aug 2017 15:48:14 +0200 Subject: Adding a scope for reduction flags. --- src/g_ltac2.ml4 | 37 ++++++++++++++++++++++++++++++++++++- src/tac2core.ml | 1 + src/tac2entries.ml | 1 + src/tac2entries.mli | 1 + src/tac2qexpr.mli | 14 ++++++++++++++ src/tac2quote.ml | 50 ++++++++++++++++++++++++++++++++++++++++++++++++++ src/tac2quote.mli | 2 ++ 7 files changed, 105 insertions(+), 1 deletion(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index c70f1e882d..ef3615db89 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -317,7 +317,7 @@ let loc_of_ne_list l = Loc.merge_opt (fst (List.hd l)) (fst (List.last l)) GEXTEND Gram GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause - q_rewriting q_clause q_dispatch q_occurrences; + q_rewriting q_clause q_dispatch q_occurrences q_strategy_flag; anti: [ [ "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] ; @@ -568,6 +568,41 @@ GEXTEND Gram q_occurrences: [ [ occs = occs -> occs ] ] ; + red_flag: + [ [ IDENT "beta" -> Loc.tag ~loc:!@loc @@ QBeta + | IDENT "iota" -> Loc.tag ~loc:!@loc @@ QIota + | IDENT "match" -> Loc.tag ~loc:!@loc @@ QMatch + | IDENT "fix" -> Loc.tag ~loc:!@loc @@ QFix + | IDENT "cofix" -> Loc.tag ~loc:!@loc @@ QCofix + | IDENT "zeta" -> Loc.tag ~loc:!@loc @@ QZeta + | IDENT "delta"; d = delta_flag -> d + ] ] + ; + refglobal: + [ [ "&"; id = Prim.ident -> QExpr (Libnames.Ident (Loc.tag ~loc:!@loc id)) + | qid = Prim.qualid -> QExpr (Libnames.Qualid qid) + | "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) + ] ] + ; + refglobals: + [ [ gl = LIST1 refglobal -> Loc.tag ~loc:!@loc gl ] ] + ; + delta_flag: + [ [ "-"; "["; idl = refglobals; "]" -> Loc.tag ~loc:!@loc @@ QDeltaBut idl + | "["; idl = refglobals; "]" -> Loc.tag ~loc:!@loc @@ QConst idl + | -> Loc.tag ~loc:!@loc @@ QDeltaBut (Loc.tag ~loc:!@loc []) + ] ] + ; + strategy_flag: + [ [ s = LIST1 red_flag -> Loc.tag ~loc:!@loc s + | d = delta_flag -> + Loc.tag ~loc:!@loc + [Loc.tag ~loc:!@loc QBeta; Loc.tag ~loc:!@loc QIota; Loc.tag ~loc:!@loc QZeta; d] + ] ] + ; + q_strategy_flag: + [ [ flag = strategy_flag -> flag ] ] + ; END (** Extension of constr syntax *) diff --git a/src/tac2core.ml b/src/tac2core.ml index 303d62a30d..196755346d 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -930,6 +930,7 @@ let () = add_expr_scope "rewriting" q_rewriting Tac2quote.of_rewriting let () = add_expr_scope "clause" q_clause Tac2quote.of_clause let () = add_expr_scope "occurrences" q_occurrences Tac2quote.of_occurrences let () = add_expr_scope "dispatch" q_dispatch Tac2quote.of_dispatch +let () = add_expr_scope "strategy" q_strategy_flag Tac2quote.of_strategy_flag let () = add_generic_scope "constr" Pcoq.Constr.constr Stdarg.wit_constr let () = add_generic_scope "open_constr" Pcoq.Constr.constr Stdarg.wit_open_constr diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 2490f6534b..a6c0e21ac5 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -33,6 +33,7 @@ let q_rewriting = Pcoq.Gram.entry_create "tactic:q_rewriting" let q_clause = Pcoq.Gram.entry_create "tactic:q_clause" let q_dispatch = Pcoq.Gram.entry_create "tactic:q_dispatch" let q_occurrences = Pcoq.Gram.entry_create "tactic:q_occurrences" +let q_strategy_flag = Pcoq.Gram.entry_create "tactic:q_strategy_flag" end (** Tactic definition *) diff --git a/src/tac2entries.mli b/src/tac2entries.mli index 667378030a..645d37a8c6 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -68,4 +68,5 @@ val q_rewriting : rewriting Pcoq.Gram.entry val q_clause : clause Pcoq.Gram.entry val q_dispatch : dispatch Pcoq.Gram.entry val q_occurrences : occurrences Pcoq.Gram.entry +val q_strategy_flag : strategy_flag Pcoq.Gram.entry end diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index 2e7fc7b44a..7c774a5c80 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -101,3 +101,17 @@ type rewriting = rewriting_r located type dispatch_r = raw_tacexpr option list * (raw_tacexpr option * raw_tacexpr option list) option type dispatch = dispatch_r located + +type red_flag_r = +| QBeta +| QIota +| QMatch +| QFix +| QCofix +| QZeta +| QConst of Libnames.reference or_anti list located +| QDeltaBut of Libnames.reference or_anti list located + +type red_flag = red_flag_r located + +type strategy_flag = red_flag list located diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 8adeb15d20..4fcbcb5424 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -236,3 +236,53 @@ let of_dispatch tacs = in let map e = of_pair default (fun l -> of_list ~loc default l) (Loc.tag ~loc e) in of_pair (fun l -> of_list ~loc default l) (fun r -> of_option ~loc map r) tacs + +let make_red_flag l = + let open Genredexpr in + let rec add_flag red = function + | [] -> red + | (_, flag) :: lf -> + let red = match flag with + | QBeta -> { red with rBeta = true } + | QMatch -> { red with rMatch = true } + | QFix -> { red with rFix = true } + | QCofix -> { red with rCofix = true } + | QZeta -> { red with rZeta = true } + | QConst (loc, l) -> + if red.rDelta then + CErrors.user_err ?loc Pp.(str + "Cannot set both constants to unfold and constants not to unfold"); + { red with rConst = red.rConst @ l } + | QDeltaBut (loc, l) -> + if red.rConst <> [] && not red.rDelta then + CErrors.user_err ?loc Pp.(str + "Cannot set both constants to unfold and constants not to unfold"); + { red with rConst = red.rConst @ l; rDelta = true } + | QIota -> + { red with rMatch = true; rFix = true; rCofix = true } + in + add_flag red lf + in + add_flag + {rBeta = false; rMatch = false; rFix = false; rCofix = false; + rZeta = false; rDelta = false; rConst = []} + l + +let of_strategy_flag (loc, flag) = + let open Genredexpr in + let loc = Option.default dummy_loc loc in + let flag = make_red_flag flag in + let of_reference ref = + let loc = Libnames.loc_of_reference ref in + inj_wit ?loc Tac2env.wit_reference ref + in + let of_ref r = of_anti of_reference r in + CTacRec (loc, [ + std_proj "rBeta", of_bool ~loc flag.rBeta; + std_proj "rMatch", of_bool ~loc flag.rMatch; + std_proj "rFix", of_bool ~loc flag.rFix; + std_proj "rCofix", of_bool ~loc flag.rCofix; + std_proj "rZeta", of_bool ~loc flag.rZeta; + std_proj "rDelta", of_bool ~loc flag.rDelta; + std_proj "rConst", of_list ~loc of_ref flag.rConst; + ]) diff --git a/src/tac2quote.mli b/src/tac2quote.mli index 456d1fa97d..730324d051 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -55,3 +55,5 @@ val of_exact_hyp : ?loc:Loc.t -> Id.t located -> raw_tacexpr (** id ↦ 'Control.refine (fun () => Control.hyp @id') *) val of_dispatch : dispatch -> raw_tacexpr + +val of_strategy_flag : strategy_flag -> raw_tacexpr -- cgit v1.2.3 From 4c964aa3ecfbb2f6aa52274545c2e27d7d11e179 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 24 Aug 2017 16:37:39 +0200 Subject: Fix typing of reference quotations. --- src/tac2core.ml | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 196755346d..b38f0b8582 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -21,6 +21,7 @@ open Proofview.Notations module Value = Tac2ffi +let std_core n = KerName.make2 Tac2env.std_prefix (Label.of_id (Id.of_string_soft n)) let coq_core n = KerName.make2 Tac2env.coq_prefix (Label.of_id (Id.of_string_soft n)) module Core = @@ -35,6 +36,7 @@ let t_constr = coq_core "constr" let t_pattern = coq_core "pattern" let t_ident = coq_core "ident" let t_option = coq_core "option" +let t_reference = std_core "reference" let c_nil = coq_core "[]" let c_cons = coq_core "::" @@ -730,7 +732,7 @@ let () = let () = let interp _ gr = return (Value.of_reference gr) in let obj = { - ml_type = t_pattern; + ml_type = t_reference; ml_interp = interp; } in define_ml_object Tac2env.wit_reference obj -- cgit v1.2.3 From 3cb2f4901ea4d79ff20b45bc4d1968ada1695d3b Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 24 Aug 2017 16:48:44 +0200 Subject: Adding notation for the remaining reduction functions. --- tests/example2.v | 25 +++++++++++++++++++++++++ theories/Notations.v | 28 ++++++++++++++++++++++------ 2 files changed, 47 insertions(+), 6 deletions(-) diff --git a/tests/example2.v b/tests/example2.v index 1856663953..af664ef163 100644 --- a/tests/example2.v +++ b/tests/example2.v @@ -157,3 +157,28 @@ intros H. split; fold (1 + 0) (1 + 0) in H. reflexivity. Qed. + +Goal 1 + 1 = 2. +Proof. +cbv [ Nat.add ]. +reflexivity. +Qed. + +Goal 1 + 1 = 2. +Proof. +let x := reference:(Nat.add) in +cbn beta iota delta [ $x ]. +reflexivity. +Qed. + +Goal 1 + 1 = 2. +Proof. +simpl beta. +reflexivity. +Qed. + +Goal 1 + 1 = 2. +Proof. +lazy. +reflexivity. +Qed. diff --git a/theories/Notations.v b/theories/Notations.v index 8c2c09a2b5..26d40fcb89 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -257,13 +257,21 @@ Ltac2 Notation "hnf" cl(opt(clause)) := Std.hnf (default_on_concl cl). Ltac2 Notation hnf := hnf. -Ltac2 Notation "vm_compute" pl(opt(seq(pattern, occurrences))) cl(opt(clause)) := - Std.vm pl (default_on_concl cl). -Ltac2 Notation vm_compute := vm_compute. +Ltac2 Notation "simpl" s(strategy) pl(opt(seq(pattern, occurrences))) cl(opt(clause)) := + Std.simpl s pl (default_on_concl cl). +Ltac2 Notation simpl := simpl. -Ltac2 Notation "native_compute" pl(opt(seq(pattern, occurrences))) cl(opt(clause)) := - Std.native pl (default_on_concl cl). -Ltac2 Notation native_compute := native_compute. +Ltac2 Notation "cbv" s(strategy) cl(opt(clause)) := + Std.cbv s (default_on_concl cl). +Ltac2 Notation cbv := cbv. + +Ltac2 Notation "cbn" s(strategy) cl(opt(clause)) := + Std.cbn s (default_on_concl cl). +Ltac2 Notation cbn := cbn. + +Ltac2 Notation "lazy" s(strategy) cl(opt(clause)) := + Std.lazy s (default_on_concl cl). +Ltac2 Notation lazy := lazy. Ltac2 fold0 pl cl := let cl := default_on_concl cl in @@ -275,6 +283,14 @@ Ltac2 Notation "fold" pl(thunk(list1(open_constr))) cl(opt(clause)) := Ltac2 Notation "pattern" pl(list1(seq(constr, occurrences), ",")) cl(opt(clause)) := Std.pattern pl (default_on_concl cl). +Ltac2 Notation "vm_compute" pl(opt(seq(pattern, occurrences))) cl(opt(clause)) := + Std.vm pl (default_on_concl cl). +Ltac2 Notation vm_compute := vm_compute. + +Ltac2 Notation "native_compute" pl(opt(seq(pattern, occurrences))) cl(opt(clause)) := + Std.native pl (default_on_concl cl). +Ltac2 Notation native_compute := native_compute. + Ltac2 rewrite0 ev rw cl tac := let tac := match tac with | None => None -- cgit v1.2.3 From 60a98c8092a0293b852712f8e21ead6e0ef1e064 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 24 Aug 2017 17:08:39 +0200 Subject: Adding a notation scope for global references. --- src/g_ltac2.ml4 | 6 +++++- src/tac2core.ml | 1 + src/tac2entries.ml | 1 + src/tac2entries.mli | 1 + src/tac2quote.ml | 14 ++++++++------ src/tac2quote.mli | 2 ++ 6 files changed, 18 insertions(+), 7 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index ef3615db89..4a7ba31373 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -317,7 +317,8 @@ let loc_of_ne_list l = Loc.merge_opt (fst (List.hd l)) (fst (List.last l)) GEXTEND Gram GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause - q_rewriting q_clause q_dispatch q_occurrences q_strategy_flag; + q_rewriting q_clause q_dispatch q_occurrences q_strategy_flag + q_reference; anti: [ [ "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] ; @@ -584,6 +585,9 @@ GEXTEND Gram | "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] ; + q_reference: + [ [ r = refglobal -> r ] ] + ; refglobals: [ [ gl = LIST1 refglobal -> Loc.tag ~loc:!@loc gl ] ] ; diff --git a/src/tac2core.ml b/src/tac2core.ml index b38f0b8582..6c38d1dfd5 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -933,6 +933,7 @@ let () = add_expr_scope "clause" q_clause Tac2quote.of_clause let () = add_expr_scope "occurrences" q_occurrences Tac2quote.of_occurrences let () = add_expr_scope "dispatch" q_dispatch Tac2quote.of_dispatch let () = add_expr_scope "strategy" q_strategy_flag Tac2quote.of_strategy_flag +let () = add_expr_scope "reference" q_reference Tac2quote.of_reference let () = add_generic_scope "constr" Pcoq.Constr.constr Stdarg.wit_constr let () = add_generic_scope "open_constr" Pcoq.Constr.constr Stdarg.wit_open_constr diff --git a/src/tac2entries.ml b/src/tac2entries.ml index a6c0e21ac5..d2b69aaf7d 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -33,6 +33,7 @@ let q_rewriting = Pcoq.Gram.entry_create "tactic:q_rewriting" let q_clause = Pcoq.Gram.entry_create "tactic:q_clause" let q_dispatch = Pcoq.Gram.entry_create "tactic:q_dispatch" let q_occurrences = Pcoq.Gram.entry_create "tactic:q_occurrences" +let q_reference = Pcoq.Gram.entry_create "tactic:q_reference" let q_strategy_flag = Pcoq.Gram.entry_create "tactic:q_strategy_flag" end diff --git a/src/tac2entries.mli b/src/tac2entries.mli index 645d37a8c6..8b92bd16f6 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -68,5 +68,6 @@ val q_rewriting : rewriting Pcoq.Gram.entry val q_clause : clause Pcoq.Gram.entry val q_dispatch : dispatch Pcoq.Gram.entry val q_occurrences : occurrences Pcoq.Gram.entry +val q_reference : Libnames.reference or_anti Pcoq.Gram.entry val q_strategy_flag : strategy_flag Pcoq.Gram.entry end diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 4fcbcb5424..9778bd18ae 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -268,15 +268,17 @@ let make_red_flag l = rZeta = false; rDelta = false; rConst = []} l +let of_reference r = + let of_ref ref = + let loc = Libnames.loc_of_reference ref in + inj_wit ?loc Tac2env.wit_reference ref + in + of_anti of_ref r + let of_strategy_flag (loc, flag) = let open Genredexpr in let loc = Option.default dummy_loc loc in let flag = make_red_flag flag in - let of_reference ref = - let loc = Libnames.loc_of_reference ref in - inj_wit ?loc Tac2env.wit_reference ref - in - let of_ref r = of_anti of_reference r in CTacRec (loc, [ std_proj "rBeta", of_bool ~loc flag.rBeta; std_proj "rMatch", of_bool ~loc flag.rMatch; @@ -284,5 +286,5 @@ let of_strategy_flag (loc, flag) = std_proj "rCofix", of_bool ~loc flag.rCofix; std_proj "rZeta", of_bool ~loc flag.rZeta; std_proj "rDelta", of_bool ~loc flag.rDelta; - std_proj "rConst", of_list ~loc of_ref flag.rConst; + std_proj "rConst", of_list ~loc of_reference flag.rConst; ]) diff --git a/src/tac2quote.mli b/src/tac2quote.mli index 730324d051..bd2303ac98 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -48,6 +48,8 @@ val of_rewriting : rewriting -> raw_tacexpr val of_occurrences : occurrences -> raw_tacexpr +val of_reference : Libnames.reference or_anti -> raw_tacexpr + val of_hyp : ?loc:Loc.t -> Id.t located -> raw_tacexpr (** id ↦ 'Control.hyp @id' *) -- cgit v1.2.3 From c515a8acb4acbe7e73121f1060ffef31d96a1436 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 24 Aug 2017 17:14:38 +0200 Subject: Adding a notation for the unfold tactic. --- tests/example2.v | 8 ++++++++ theories/Notations.v | 3 +++ 2 files changed, 11 insertions(+) diff --git a/tests/example2.v b/tests/example2.v index af664ef163..6b26b78022 100644 --- a/tests/example2.v +++ b/tests/example2.v @@ -182,3 +182,11 @@ Proof. lazy. reflexivity. Qed. + +Goal let x := 1 + 1 - 1 in x = x. +Proof. +intros x. +unfold &x at 1. +let x := reference:(Nat.sub) in unfold Nat.add, $x in x. +reflexivity. +Qed. diff --git a/theories/Notations.v b/theories/Notations.v index 26d40fcb89..62d5d65d7c 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -273,6 +273,9 @@ Ltac2 Notation "lazy" s(strategy) cl(opt(clause)) := Std.lazy s (default_on_concl cl). Ltac2 Notation lazy := lazy. +Ltac2 Notation "unfold" pl(list1(seq(reference, occurrences), ",")) cl(opt(clause)) := + Std.unfold pl (default_on_concl cl). + Ltac2 fold0 pl cl := let cl := default_on_concl cl in Control.enter (fun () => Control.with_holes pl (fun pl => Std.fold pl cl)). -- cgit v1.2.3 From a3bef204ec2840b879c37b0b3ba43574a6550647 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 24 Aug 2017 17:32:46 +0200 Subject: Rely on quoting for lists instead of hardwiring it in the AST. --- src/g_ltac2.ml4 | 3 ++- src/tac2core.ml | 20 ++++---------------- src/tac2expr.mli | 1 - src/tac2intern.ml | 15 --------------- 4 files changed, 6 insertions(+), 33 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 4a7ba31373..7406a45207 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -125,7 +125,8 @@ GEXTEND Gram | "("; a = SELF; ":"; t = tac2type; ")" -> CTacCnv (!@loc, a, t) | "()" -> CTacCst (!@loc, AbsKn (Tuple 0)) | "("; ")" -> CTacCst (!@loc, AbsKn (Tuple 0)) - | "["; a = LIST0 tac2expr LEVEL "5" SEP ";"; "]" -> CTacLst (Loc.tag ~loc:!@loc a) + | "["; a = LIST0 tac2expr LEVEL "5" SEP ";"; "]" -> + Tac2quote.of_list ~loc:!@loc (fun x -> x) a | "{"; a = tac2rec_fieldexprs; "}" -> CTacRec (!@loc, a) | a = tactic_atom -> a ] ] diff --git a/src/tac2core.ml b/src/tac2core.ml index 6c38d1dfd5..8d0f640209 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -827,19 +827,13 @@ let () = add_scope "list0" begin function | [tok] -> let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in let scope = Extend.Alist0 scope in - let act l = - let l = List.map act l in - CTacLst (None, l) - in + let act l = Tac2quote.of_list act l in Tac2entries.ScopeRule (scope, act) | [tok; SexprStr (_, str)] -> let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in let sep = Extend.Atoken (CLexer.terminal str) in let scope = Extend.Alist0sep (scope, sep) in - let act l = - let l = List.map act l in - CTacLst (None, l) - in + let act l = Tac2quote.of_list act l in Tac2entries.ScopeRule (scope, act) | _ -> scope_fail () end @@ -848,19 +842,13 @@ let () = add_scope "list1" begin function | [tok] -> let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in let scope = Extend.Alist1 scope in - let act l = - let l = List.map act l in - CTacLst (None, l) - in + let act l = Tac2quote.of_list act l in Tac2entries.ScopeRule (scope, act) | [tok; SexprStr (_, str)] -> let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in let sep = Extend.Atoken (CLexer.terminal str) in let scope = Extend.Alist1sep (scope, sep) in - let act l = - let l = List.map act l in - CTacLst (None, l) - in + let act l = Tac2quote.of_list act l in Tac2entries.ScopeRule (scope, act) | _ -> scope_fail () end diff --git a/src/tac2expr.mli b/src/tac2expr.mli index 7efb85cbb0..76fb50181f 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -99,7 +99,6 @@ type raw_tacexpr = | CTacApp of Loc.t * raw_tacexpr * raw_tacexpr list | CTacLet of Loc.t * rec_flag * (raw_patexpr * raw_typexpr option * raw_tacexpr) list * raw_tacexpr | CTacArr of raw_tacexpr list located -| CTacLst of raw_tacexpr list located | CTacCnv of Loc.t * raw_tacexpr * raw_typexpr | CTacSeq of Loc.t * raw_tacexpr * raw_tacexpr | CTacCse of Loc.t * raw_tacexpr * raw_taccase list diff --git a/src/tac2intern.ml b/src/tac2intern.ml index b62a574a6c..865935c52d 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -195,7 +195,6 @@ let loc_of_tacexpr = function | CTacApp (loc, _, _) -> loc | CTacLet (loc, _, _, _) -> loc | CTacArr (loc, _) -> Option.default dummy_loc loc -| CTacLst (loc, _) -> Option.default dummy_loc loc | CTacCnv (loc, _, _) -> loc | CTacSeq (loc, _, _) -> loc | CTacCse (loc, _, _) -> loc @@ -735,14 +734,6 @@ let rec intern_rec env = function let fold e el = intern_rec_with_constraint env e t0 :: el in let el = e0 :: List.fold_right fold el [] in (GTacArr el, GTypRef (Other t_array, [t0])) -| CTacLst (loc, []) -> - let id = fresh_id env in - (c_nil, GTypRef (Other t_list, [GTypVar id])) -| CTacLst (loc, e0 :: el) -> - let (e0, t0) = intern_rec env e0 in - let fold e el = c_cons (intern_rec_with_constraint env e t0) el in - let el = c_cons e0 (List.fold_right fold el c_nil) in - (el, GTypRef (Other t_list, [t0])) | CTacCnv (loc, e, tc) -> let (e, t) = intern_rec env e in let tc = intern_type env tc in @@ -1229,9 +1220,6 @@ let rec globalize ids e = match e with | CTacArr (loc, el) -> let el = List.map (fun e -> globalize ids e) el in CTacArr (loc, el) -| CTacLst (loc, el) -> - let el = List.map (fun e -> globalize ids e) el in - CTacLst (loc, el) | CTacCnv (loc, e, t) -> let e = globalize ids e in CTacCnv (loc, e, t) @@ -1446,9 +1434,6 @@ let rec subst_rawexpr subst t = match t with | CTacArr (loc, el) -> let el' = List.smartmap (fun e -> subst_rawexpr subst e) el in if el' == el then t else CTacArr (loc, el') -| CTacLst (loc, el) -> - let el' = List.smartmap (fun e -> subst_rawexpr subst e) el in - if el' == el then t else CTacLst (loc, el') | CTacCnv (loc, e, c) -> let e' = subst_rawexpr subst e in let c' = subst_rawtype subst c in -- cgit v1.2.3 From 6a5558405f801c466a51f32080c8dbb893a2170d Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 24 Aug 2017 17:34:55 +0200 Subject: Removing dead code about arrays in the AST. --- src/tac2expr.mli | 1 - src/tac2intern.ml | 15 --------------- 2 files changed, 16 deletions(-) diff --git a/src/tac2expr.mli b/src/tac2expr.mli index 76fb50181f..281ed6c81e 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -98,7 +98,6 @@ type raw_tacexpr = | CTacFun of Loc.t * (raw_patexpr * raw_typexpr option) list * raw_tacexpr | CTacApp of Loc.t * raw_tacexpr * raw_tacexpr list | CTacLet of Loc.t * rec_flag * (raw_patexpr * raw_typexpr option * raw_tacexpr) list * raw_tacexpr -| CTacArr of raw_tacexpr list located | CTacCnv of Loc.t * raw_tacexpr * raw_typexpr | CTacSeq of Loc.t * raw_tacexpr * raw_tacexpr | CTacCse of Loc.t * raw_tacexpr * raw_taccase list diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 865935c52d..765be92103 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -194,7 +194,6 @@ let loc_of_tacexpr = function | CTacFun (loc, _, _) -> loc | CTacApp (loc, _, _) -> loc | CTacLet (loc, _, _, _) -> loc -| CTacArr (loc, _) -> Option.default dummy_loc loc | CTacCnv (loc, _, _) -> loc | CTacSeq (loc, _, _) -> loc | CTacCse (loc, _, _) -> loc @@ -726,14 +725,6 @@ let rec intern_rec env = function let ids = List.fold_left fold Id.Set.empty el in if is_rec then intern_let_rec env loc ids el e else intern_let env loc ids el e -| CTacArr (loc, []) -> - let id = fresh_id env in - (GTacArr [], GTypRef (Other t_int, [GTypVar id])) -| CTacArr (loc, e0 :: el) -> - let (e0, t0) = intern_rec env e0 in - let fold e el = intern_rec_with_constraint env e t0 :: el in - let el = e0 :: List.fold_right fold el [] in - (GTacArr el, GTypRef (Other t_array, [t0])) | CTacCnv (loc, e, tc) -> let (e, t) = intern_rec env e in let tc = intern_type env tc in @@ -1217,9 +1208,6 @@ let rec globalize ids e = match e with in let bnd = List.map map bnd in CTacLet (loc, isrec, bnd, e) -| CTacArr (loc, el) -> - let el = List.map (fun e -> globalize ids e) el in - CTacArr (loc, el) | CTacCnv (loc, e, t) -> let e = globalize ids e in CTacCnv (loc, e, t) @@ -1431,9 +1419,6 @@ let rec subst_rawexpr subst t = match t with let bnd' = List.smartmap map bnd in let e' = subst_rawexpr subst e in if bnd' == bnd && e' == e then t else CTacLet (loc, isrec, bnd', e') -| CTacArr (loc, el) -> - let el' = List.smartmap (fun e -> subst_rawexpr subst e) el in - if el' == el then t else CTacArr (loc, el') | CTacCnv (loc, e, c) -> let e' = subst_rawexpr subst e in let c' = subst_rawtype subst c in -- cgit v1.2.3 From 7cd041b42588e6d9ff0e5ea127960585666c4b07 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 24 Aug 2017 17:54:28 +0200 Subject: Documentation about the transition from Ltac1. --- doc/ltac2.md | 106 +++++++++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 106 insertions(+) diff --git a/doc/ltac2.md b/doc/ltac2.md index 6c4912c8f3..55780a7712 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -690,6 +690,112 @@ foo 0 ↦ (fun x => x ()) (fun _ => 0) Note that abbreviations are not typechecked at all, and may result in typing errors after expansion. +# Transition from Ltac1 + +Owing to the use of a bunch of notations, the transition shouldn't be +atrociously horrible and shockingly painful up to the point you want to retire +in the Ariège mountains, living off the land and insulting careless bypassers in +proto-georgian. + +That said, we do *not* guarantee you it is going to be a blissful walk either. +Hopefully, owing to the fact Ltac2 is typed, the interactive dialogue with Coq +will help you. + +We list the major changes and the transition strategies hereafter. + +## Syntax changes + +Due to conflicts, a few syntactic rules have changed. + +- The dispatch tactical `tac; [foo|bar]` is now written `tac > [foo|bar]`. +- Levels of a few operators have been revised. Some tacticals now parse as if + they were a normal function, i.e. one has to put parentheses around the + argument when it is complex, e.g an abstraction. List of affected tacticals: + `try`, `repeat`, `do`, `once`, `progress`, `time`, `abstract`. +- `idtac` is no more. Either use `()` if you expect nothing to happen, + `(fun () => ())` if you want a thunk (see next section), or use printing + primitives from the `Message` module if you wand to display something. + +## Tactic delay + +Tactics are not magically delayed anymore, neither as functions nor as +arguments. It is your responsibility to thunk them beforehand and apply them +at the call site. + +A typical example of a delayed function: +``` +Ltac foo := blah. +``` +becomes +``` +Ltac2 foo () := blah. +``` + +All subsequent calls to `foo` must be applied to perform the same effect as +before. + +Likewise, for arguments: +``` +Ltac bar tac := tac; tac; tac. +``` +becomes +``` +Ltac2 bar tac := tac (); tac (); tac (). +``` + +We recommend the use of syntactic notations to ease the transition. For +instance, the first example can alternatively written as: +``` +Ltac2 foo0 () := blah. +Ltac2 Notation foo := foo0 (). +``` + +This allows to keep the subsequent calls to the tactic as-is, as the +expression `foo` will be implicitly expanded everywhere into `foo0 ()`. Such +a trick also works for arguments, as arguments of syntactic notations are +implicitly thunked. The second example could thus be written as follows. + +``` +Ltac2 bar0 tac := tac (); tac (); tac (). +Ltac2 Notation bar := bar0. +``` + +## Variable binding + +Ltac1 relies on a crazy amount of dynamic trickery to be able to tell apart +bound variables from terms, hypotheses and whatnot. There is no such thing in +Ltac2, as variables are recognized statically and other constructions do not +live in the same syntactic world. Due to the abuse of quotations, it can +sometimes be complicated to know what a mere identifier represent in a tactic +expression. We recommend tracking the context and letting the compiler spit +typing errors to understand what is going on. + +We list below the typical changes one has to perform depending on the static +errors produced by the typechecker. + +### In Ltac expressions + +- `Unbound value X`, `Unbound constructor X`: + * if `X` is meant to be a term from the current stactic environment, replace + the problematic use by `'X`. + * if `X` is meant to be a hypothesis from the goal context, replace the + problematic use by `&X`. + +### In quotations + +- `The reference X was not found in the current environment`: + * if `X` is meant to be a tactic expression bound by a Ltac2 let or function, + replace the problematic use by `$X`. + * if `X` is meant to be a hypothesis from the goal context, replace the + problematic use by `&X`. + +## Exception catching + +Ltac2 features a proper exception-catching mechanism. For this reason, the +Ltac1 mechanism relying on `fail` taking integers and tacticals decreasing it +has been removed. Now exceptions are preserved by all tacticals, and it is +your duty to catch it and reraise it depending on your use. + # TODO - Implement deep pattern-matching. -- cgit v1.2.3 From 72f47973b860c8074aa976759ee1adce993dac49 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 24 Aug 2017 18:42:21 +0200 Subject: Fix the semantics of fail, as it should enter the goal. --- theories/Notations.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/Notations.v b/theories/Notations.v index 62d5d65d7c..d2c7059985 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -27,7 +27,7 @@ match Control.case t with Control.plus (fun _ => s x) (fun e => s (k e)) end. -Ltac2 fail0 (_ : unit) := Control.zero Tactic_failure. +Ltac2 fail0 (_ : unit) := Control.enter (fun _ => Control.zero Tactic_failure). Ltac2 Notation fail := fail0 (). -- cgit v1.2.3 From c3d83f8437022986593df45c3c3920c8356d9f84 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 25 Aug 2017 14:51:20 +0200 Subject: Lookahead to cheat the constr parser in order to parse "& IDENT". --- src/g_ltac2.ml4 | 14 ++++++++++++-- 1 file changed, 12 insertions(+), 2 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 7406a45207..3d873c7369 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -54,6 +54,16 @@ let test_lpar_id_rpar = | _ -> err ()) | _ -> err ()) +let test_ampersand_ident = + Gram.Entry.of_parser "test_ampersand_ident" + (fun strm -> + match stream_nth 0 strm with + | KEYWORD "&" -> + (match stream_nth 1 strm with + | IDENT _ -> () + | _ -> err ()) + | _ -> err ()) + let tac2expr = Tac2entries.Pltac.tac2expr let tac2type = Gram.entry_create "tactic:tac2type" let tac2def_val = Gram.entry_create "tactic:tac2def_val" @@ -617,8 +627,8 @@ GEXTEND Gram [ [ IDENT "ltac2"; ":"; "("; tac = tac2expr; ")" -> let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in CAst.make ~loc:!@loc (CHole (None, IntroAnonymous, Some arg)) - | "&"; id = [ id = Prim.ident -> Loc.tag ~loc:!@loc id ] -> - let tac = Tac2quote.of_exact_hyp ~loc:!@loc id in + | test_ampersand_ident; "&"; id = Prim.ident -> + let tac = Tac2quote.of_exact_hyp ~loc:!@loc (Loc.tag ~loc:!@loc id) in let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in CAst.make ~loc:!@loc (CHole (None, IntroAnonymous, Some arg)) ] ] -- cgit v1.2.3 From 3ef5f35b0d7ec6f56f68e4319d6ec85bebaa19b8 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 25 Aug 2017 14:57:07 +0200 Subject: Parse specifically idents as destruction arguments. This is for backward compatibility. --- src/g_ltac2.ml4 | 1 + 1 file changed, 1 insertion(+) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 3d873c7369..c1025ceba5 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -459,6 +459,7 @@ GEXTEND Gram ; destruction_arg: [ [ n = lnatural -> QElimOnAnonHyp n + | id = lident -> QElimOnIdent id | c = constr_with_bindings -> QElimOnConstr c ] ] ; -- cgit v1.2.3 From c41f5d406f627e94363b4549ef268ffa33e7b681 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 25 Aug 2017 15:01:02 +0200 Subject: Respect the default goal selector in toplevel invocations. --- src/tac2entries.ml | 3 ++- tests/example2.v | 2 ++ 2 files changed, 4 insertions(+), 1 deletion(-) diff --git a/src/tac2entries.ml b/src/tac2entries.ml index d2b69aaf7d..73086c406e 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -700,7 +700,8 @@ let print_ltac ref = let solve default tac = let status = Proof_global.with_current_proof begin fun etac p -> let with_end_tac = if default then Some etac else None in - let (p, status) = Pfedit.solve SelectAll None tac ?with_end_tac p in + let g = Proof_bullet.get_default_goal_selector () in + let (p, status) = Pfedit.solve g None tac ?with_end_tac p in (* in case a strict subtree was completed, go back to the top of the prooftree *) let p = Proof.maximal_unfocus Vernacentries.command_focus p in diff --git a/tests/example2.v b/tests/example2.v index 6b26b78022..95485305dc 100644 --- a/tests/example2.v +++ b/tests/example2.v @@ -2,6 +2,8 @@ Require Import Ltac2.Ltac2. Import Ltac2.Notations. +Set Default Goal Selector "all". + Goal exists n, n = 0. Proof. split with (x := 0). -- cgit v1.2.3 From 47eb0278a3cdf93129b1742e314681d65bd6475a Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 25 Aug 2017 14:11:48 +0200 Subject: More bindings to primitive tactics. --- src/tac2stdlib.ml | 26 ++++++++++++++++++++++++++ src/tac2tactics.ml | 49 ++++++++++++++++++++++++++++++++++++++++++++++++- src/tac2tactics.mli | 10 +++++++++- theories/Notations.v | 33 ++++++++++++++++++++++++++++++++- theories/Std.v | 10 ++++++++++ 5 files changed, 125 insertions(+), 3 deletions(-) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index eccaf95ab3..d88402cbf2 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -481,10 +481,36 @@ end (** Tactics from extratactics *) +let () = define_prim2 "tac_discriminate" begin fun ev arg -> + let ev = Value.to_bool ev in + let arg = Value.to_option (fun arg -> None, to_destruction_arg arg) arg in + Tac2tactics.discriminate ev arg +end + +let () = define_prim3 "tac_injection" begin fun ev ipat arg -> + let ev = Value.to_bool ev in + let ipat = Value.to_option to_intro_patterns ipat in + let arg = Value.to_option (fun arg -> None, to_destruction_arg arg) arg in + Tac2tactics.injection ev ipat arg +end + let () = define_prim1 "tac_absurd" begin fun c -> Contradiction.absurd (Value.to_constr c) end +let () = define_prim1 "tac_contradiction" begin fun c -> + let c = Value.to_option to_constr_with_bindings c in + Contradiction.contradiction c +end + +let () = define_prim4 "tac_autorewrite" begin fun all by ids cl -> + let all = Value.to_bool all in + let by = Value.to_option (fun tac -> Proofview.tclIGNORE (thaw tac)) by in + let ids = Value.to_list Value.to_ident ids in + let cl = to_clause cl in + Tac2tactics.autorewrite ~all by ids cl +end + let () = define_prim1 "tac_subst" begin fun ids -> let ids = Value.to_list Value.to_ident ids in Equality.subst ids diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index 2f4965783f..25a00fdc2e 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -16,6 +16,15 @@ open Genredexpr open Proofview open Proofview.Notations +type destruction_arg = EConstr.constr with_bindings tactic Misctypes.destruction_arg + +let tactic_infer_flags with_evar = { + Pretyping.use_typeclasses = true; + Pretyping.solve_unification_constraints = true; + Pretyping.use_hook = None; + Pretyping.fail_evar = not with_evar; + Pretyping.expand_evars = true } + (** FIXME: export a better interface in Tactics *) let delayed_of_tactic tac env sigma = let _, pv = Proofview.init sigma [] in @@ -30,7 +39,7 @@ let apply adv ev cb cl = | Some (id, cl) -> Tactics.apply_delayed_in adv ev id cb cl type induction_clause = - EConstr.constr with_bindings tactic destruction_arg * + destruction_arg * intro_pattern_naming option * or_and_intro_pattern option * Locus.clause option @@ -112,3 +121,41 @@ let vm where cl = let native where cl = let where = Option.map map_pattern_with_occs where in Tactics.reduce (CbvNative where) cl + +let on_destruction_arg tac ev arg = + Proofview.Goal.enter begin fun gl -> + match arg with + | None -> tac ev None + | Some (clear, arg) -> + let arg = match arg with + | ElimOnConstr c -> + let env = Proofview.Goal.env gl in + Proofview.tclEVARMAP >>= fun sigma -> + c >>= fun (c, lbind) -> + Proofview.tclEVARMAP >>= fun sigma' -> + let flags = tactic_infer_flags ev in + let (sigma', c) = Unification.finish_evar_resolution ~flags env sigma' (sigma, c) in + Proofview.tclUNIT (Some sigma', ElimOnConstr (c, lbind)) + | ElimOnIdent id -> Proofview.tclUNIT (None, ElimOnIdent id) + | ElimOnAnonHyp n -> Proofview.tclUNIT (None, ElimOnAnonHyp n) + in + arg >>= fun (sigma', arg) -> + let arg = Some (clear, arg) in + match sigma' with + | None -> tac ev arg + | Some sigma' -> + Tacticals.New.tclWITHHOLES ev (tac ev arg) sigma' + end + +let discriminate ev arg = on_destruction_arg Equality.discr_tac ev arg + +let injection ev ipat arg = + let tac ev arg = Equality.injClause ipat ev arg in + on_destruction_arg tac ev arg + +let autorewrite ~all by ids cl = + let conds = if all then Some Equality.AllMatches else None in + let ids = List.map Id.to_string ids in + match by with + | None -> Autorewrite.auto_multi_rewrite ?conds ids cl + | Some by -> Autorewrite.auto_multi_rewrite_with ?conds by ids cl diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli index d835f768a1..8939d2a9dd 100644 --- a/src/tac2tactics.mli +++ b/src/tac2tactics.mli @@ -14,6 +14,8 @@ open Misctypes open Tactypes open Proofview +type destruction_arg = EConstr.constr with_bindings tactic Misctypes.destruction_arg + (** Local reimplementations of tactics variants from Coq *) val apply : advanced_flag -> evars_flag -> @@ -21,7 +23,7 @@ val apply : advanced_flag -> evars_flag -> (Id.t * intro_pattern option) option -> unit tactic type induction_clause = - EConstr.constr with_bindings tactic destruction_arg * + destruction_arg * intro_pattern_naming option * or_and_intro_pattern option * clause option @@ -51,3 +53,9 @@ val unfold : (global_reference * occurrences_expr) list -> clause -> unit tactic val vm : (Pattern.constr_pattern * occurrences_expr) option -> clause -> unit tactic val native : (Pattern.constr_pattern * occurrences_expr) option -> clause -> unit tactic + +val discriminate : evars_flag -> destruction_arg option -> unit tactic + +val injection : evars_flag -> intro_pattern list option -> destruction_arg option -> unit tactic + +val autorewrite : all:bool -> unit tactic option -> Id.t list -> clause -> unit tactic diff --git a/theories/Notations.v b/theories/Notations.v index d2c7059985..2d52904faf 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -130,6 +130,23 @@ Ltac2 Notation split := split. Ltac2 Notation "esplit" bnd(thunk(bindings)) := split0 true bnd. Ltac2 Notation esplit := esplit. +Ltac2 exists0 ev bnds := match bnds with +| [] => split0 ev (fun () => Std.NoBindings) +| _ => + let rec aux bnds := match bnds with + | [] => () + | bnd :: bnds => split0 ev bnd; aux bnds + end in + aux bnds +end. + +(* +Ltac2 Notation "exists" bnd(list0(thunk(bindings), ",")) := exists0 false bnd. + +Ltac2 Notation "eexists" bnd(list0(thunk(bindings), ",")) := exists0 true bnd. +Ltac2 Notation eexists := eexists. +*) + Ltac2 left0 ev bnd := enter_h ev Std.left bnd. Ltac2 Notation "left" bnd(thunk(bindings)) := left0 false bnd. @@ -316,7 +333,7 @@ Ltac2 Notation "erewrite" tac(opt(seq("by", thunk(tactic)))) := rewrite0 true rw cl tac. -(** Other base tactics *) +(** coretactics *) Ltac2 Notation reflexivity := Std.reflexivity (). @@ -329,3 +346,17 @@ Ltac2 Notation admit := Std.admit (). Ltac2 Notation clear := Std.keep []. Ltac2 Notation refine := Control.refine. + +(** extratactics *) + +Ltac2 absurd0 c := Control.enter (fun _ => Std.absurd (c ())). + +Ltac2 Notation absurd := absurd0. + +Ltac2 subst0 ids := match ids with +| [] => Std.subst_all () +| _ => Std.subst ids +end. + +Ltac2 Notation "subst" ids(list0(ident)) := subst0 ids. +Ltac2 Notation subst := subst. diff --git a/theories/Std.v b/theories/Std.v index dd81835c40..f380c10af8 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -184,7 +184,17 @@ Ltac2 @ external exact_no_check : constr -> unit := "ltac2" "tac_exactnocheck". Ltac2 @ external vm_cast_no_check : constr -> unit := "ltac2" "tac_vmcastnocheck". Ltac2 @ external native_cast_no_check : constr -> unit := "ltac2" "tac_nativecastnocheck". +(** coretactics *) + +(** extratactics *) + +Ltac2 @ external discriminate : evar_flag -> destruction_arg option -> unit := "ltac2" "tac_discriminate". +Ltac2 @ external injection : evar_flag -> intro_pattern list option -> destruction_arg option -> unit := "ltac2" "tac_injection". + Ltac2 @ external absurd : constr -> unit := "ltac2" "tac_absurd". +Ltac2 @ external contradiction : constr_with_bindings option -> unit := "ltac2" "tac_contradiction". + +Ltac2 @ external autorewrite : bool -> (unit -> unit) option -> ident list -> clause -> unit := "ltac2" "tac_autorewrite". Ltac2 @ external subst : ident list -> unit := "ltac2" "tac_subst". Ltac2 @ external subst_all : unit -> unit := "ltac2" "tac_substall". -- cgit v1.2.3 From 6875b016b0a502b03296e5f97f26cf0f6699a7aa Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 25 Aug 2017 17:16:01 +0200 Subject: Do not return STRING scopes in the tuple produced by "seq" scopes. --- doc/ltac2.md | 6 ++++-- src/tac2core.ml | 13 ++++++++----- src/tac2quote.ml | 7 ++++++- src/tac2quote.mli | 2 ++ theories/Notations.v | 33 +++------------------------------ 5 files changed, 23 insertions(+), 38 deletions(-) diff --git a/doc/ltac2.md b/doc/ltac2.md index 55780a7712..a645331e2d 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -615,8 +615,10 @@ The following scopes are built-in. + parses the string *s* as a keyword, if it is already a keyword, otherwise as an IDENT. Returns `()`. - seq(*scope₁*, ..., *scopeₙ*): - + parses *scope₁*, ..., *scopeₙ* in this order, and produces a n-tuple made - out of the parsed values in the same order. It is forbidden for the various + + parses *scope₁*, ..., *scopeₙ* in this order, and produces a tuple made + out of the parsed values in the same order. As an optimization, all + subscopes of the form STRING are left out of the returned tuple, instead + of returning a useless unit value. It is forbidden for the various subscopes to refer to the global entry using self of next. For now there is no way to declare new scopes from Ltac2 side, but this is diff --git a/src/tac2core.ml b/src/tac2core.ml index 8d0f640209..7a8f3ceb44 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -963,13 +963,12 @@ let rec generalize_symbol : type _ converter = | CvNil : (Loc.t -> raw_tacexpr) converter -| CvCns : 'act converter * ('a -> raw_tacexpr) -> ('a -> 'act) converter +| CvCns : 'act converter * ('a -> raw_tacexpr) option -> ('a -> 'act) converter let rec apply : type a. a converter -> raw_tacexpr list -> a = function -| CvNil -> fun accu loc -> - let cst = CTacCst (loc, AbsKn (Tuple (List.length accu))) in - CTacApp (loc, cst, accu) -| CvCns (c, f) -> fun accu x -> apply c (f x :: accu) +| CvNil -> fun accu loc -> Tac2quote.of_tuple ~loc accu +| CvCns (c, None) -> fun accu x -> apply c accu +| CvCns (c, Some f) -> fun accu x -> apply c (f x :: accu) type seqrule = | Seqrule : ('act, Loc.t -> raw_tacexpr) norec_rule * 'act converter -> seqrule @@ -983,6 +982,10 @@ let rec make_seq_rule = function let scope = generalize_symbol scope in let Seqrule (r, c) = make_seq_rule rem in let r = { norec_rule = Next (r.norec_rule, scope.any_symbol) } in + let f = match tok with + | SexprStr _ -> None (** Leave out mere strings *) + | _ -> Some f + in Seqrule (r, CvCns (c, f)) let () = add_scope "seq" begin fun toks -> diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 9778bd18ae..dbd2fd0529 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -46,7 +46,12 @@ let of_pair f g (loc, (e1, e2)) = let loc = Option.default dummy_loc loc in CTacApp (loc, CTacCst (loc, AbsKn (Tuple 2)), [f e1; g e2]) -let of_tuple ?loc el = +let of_tuple ?loc el = match el with +| [] -> + let loc = Option.default dummy_loc loc in + CTacCst (loc, AbsKn (Tuple 0)) +| [e] -> e +| el -> let loc = Option.default dummy_loc loc in let len = List.length el in CTacApp (loc, CTacCst (loc, AbsKn (Tuple len)), el) diff --git a/src/tac2quote.mli b/src/tac2quote.mli index bd2303ac98..7f3d9dce6e 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -24,6 +24,8 @@ val of_int : int located -> raw_tacexpr val of_pair : ('a -> raw_tacexpr) -> ('b -> raw_tacexpr) -> ('a * 'b) located -> raw_tacexpr +val of_tuple : ?loc:Loc.t -> raw_tacexpr list -> raw_tacexpr + val of_variable : Id.t located -> raw_tacexpr val of_ident : Id.t located -> raw_tacexpr diff --git a/theories/Notations.v b/theories/Notations.v index 2d52904faf..bb27a34627 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -175,14 +175,7 @@ Ltac2 Notation econstructor := econstructor. Ltac2 Notation "econstructor" n(tactic) bnd(thunk(bindings)) := constructor0 true n bnd. Ltac2 elim0 ev c bnd use := - let f ev ((c, bnd, use)) := - let use := match use with - | None => None - | Some u => - let ((_, c, wth)) := u in Some (c, wth) - end in - Std.elim ev (c, bnd) use - in + let f ev ((c, bnd, use)) := Std.elim ev (c, bnd) use in enter_h ev f (fun () => c (), bnd (), use ()). Ltac2 Notation "elim" c(thunk(constr)) bnd(thunk(bindings)) @@ -219,14 +212,7 @@ Ltac2 Notation "apply" apply0 true false cb cl. Ltac2 induction0 ev ic use := - let f ev use := - let use := match use with - | None => None - | Some u => - let ((_, c, wth)) := u in Some (c, wth) - end in - Std.induction ev ic use - in + let f ev use := Std.induction ev ic use in enter_h ev f use. Ltac2 Notation "induction" @@ -240,14 +226,7 @@ Ltac2 Notation "einduction" induction0 true ic use. Ltac2 destruct0 ev ic use := - let f ev use := - let use := match use with - | None => None - | Some u => - let ((_, c, wth)) := u in Some (c, wth) - end in - Std.destruct ev ic use - in + let f ev use := Std.destruct ev ic use in enter_h ev f use. Ltac2 Notation "destruct" @@ -312,12 +291,6 @@ Ltac2 Notation "native_compute" pl(opt(seq(pattern, occurrences))) cl(opt(clause Ltac2 Notation native_compute := native_compute. Ltac2 rewrite0 ev rw cl tac := - let tac := match tac with - | None => None - | Some p => - let ((_, tac)) := p in - Some tac - end in let cl := default_on_concl cl in Std.rewrite ev rw cl tac. -- cgit v1.2.3 From 029dd0fcfc5641b689356c467e2f0fb1d3fa178c Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 25 Aug 2017 18:02:49 +0200 Subject: Renaming the bindings scope into with_bindings. --- src/g_ltac2.ml4 | 4 ++-- src/tac2core.ml | 2 +- src/tac2entries.ml | 2 +- src/tac2entries.mli | 2 +- theories/Notations.v | 40 ++++++++++++++++++++-------------------- 5 files changed, 25 insertions(+), 25 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index c1025ceba5..95fcf79000 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -327,7 +327,7 @@ open Tac2entries.Pltac let loc_of_ne_list l = Loc.merge_opt (fst (List.hd l)) (fst (List.last l)) GEXTEND Gram - GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause + GLOBAL: q_ident q_with_bindings q_intropattern q_intropatterns q_induction_clause q_rewriting q_clause q_dispatch q_occurrences q_strategy_flag q_reference; anti: @@ -365,7 +365,7 @@ GEXTEND Gram Loc.tag ~loc:!@loc @@ QImplicitBindings bl ] ] ; - q_bindings: + q_with_bindings: [ [ bl = with_bindings -> bl ] ] ; intropatterns: diff --git a/src/tac2core.ml b/src/tac2core.ml index 7a8f3ceb44..342e8f51e8 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -912,7 +912,7 @@ let add_expr_scope name entry f = end let () = add_expr_scope "ident" q_ident (fun id -> Tac2quote.of_anti Tac2quote.of_ident id) -let () = add_expr_scope "bindings" q_bindings Tac2quote.of_bindings +let () = add_expr_scope "with_bindings" q_with_bindings Tac2quote.of_bindings let () = add_expr_scope "intropattern" q_intropattern Tac2quote.of_intro_pattern let () = add_expr_scope "intropatterns" q_intropatterns Tac2quote.of_intro_patterns let () = add_expr_scope "induction_clause" q_induction_clause Tac2quote.of_induction_clause diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 73086c406e..6c8c2348fe 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -25,7 +25,7 @@ struct let tac2expr = Pcoq.Gram.entry_create "tactic:tac2expr" let q_ident = Pcoq.Gram.entry_create "tactic:q_ident" -let q_bindings = Pcoq.Gram.entry_create "tactic:q_bindings" +let q_with_bindings = Pcoq.Gram.entry_create "tactic:q_with_bindings" let q_intropattern = Pcoq.Gram.entry_create "tactic:q_intropattern" let q_intropatterns = Pcoq.Gram.entry_create "tactic:q_intropatterns" let q_induction_clause = Pcoq.Gram.entry_create "tactic:q_induction_clause" diff --git a/src/tac2entries.mli b/src/tac2entries.mli index 8b92bd16f6..d22ae7a953 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -60,7 +60,7 @@ val tac2expr : raw_tacexpr Pcoq.Gram.entry open Tac2qexpr val q_ident : Id.t located or_anti Pcoq.Gram.entry -val q_bindings : bindings Pcoq.Gram.entry +val q_with_bindings : bindings Pcoq.Gram.entry val q_intropattern : intro_pattern Pcoq.Gram.entry val q_intropatterns : intro_pattern list located Pcoq.Gram.entry val q_induction_clause : induction_clause Pcoq.Gram.entry diff --git a/theories/Notations.v b/theories/Notations.v index bb27a34627..8345344d94 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -124,10 +124,10 @@ Ltac2 Notation eintros := eintros. Ltac2 split0 ev bnd := enter_h ev Std.split bnd. -Ltac2 Notation "split" bnd(thunk(bindings)) := split0 false bnd. +Ltac2 Notation "split" bnd(thunk(with_bindings)) := split0 false bnd. Ltac2 Notation split := split. -Ltac2 Notation "esplit" bnd(thunk(bindings)) := split0 true bnd. +Ltac2 Notation "esplit" bnd(thunk(with_bindings)) := split0 true bnd. Ltac2 Notation esplit := esplit. Ltac2 exists0 ev bnds := match bnds with @@ -141,26 +141,26 @@ Ltac2 exists0 ev bnds := match bnds with end. (* -Ltac2 Notation "exists" bnd(list0(thunk(bindings), ",")) := exists0 false bnd. +Ltac2 Notation "exists" bnd(list0(thunk(with_bindings), ",")) := exists0 false bnd. -Ltac2 Notation "eexists" bnd(list0(thunk(bindings), ",")) := exists0 true bnd. +Ltac2 Notation "eexists" bnd(list0(thunk(with_bindings), ",")) := exists0 true bnd. Ltac2 Notation eexists := eexists. *) Ltac2 left0 ev bnd := enter_h ev Std.left bnd. -Ltac2 Notation "left" bnd(thunk(bindings)) := left0 false bnd. +Ltac2 Notation "left" bnd(thunk(with_bindings)) := left0 false bnd. Ltac2 Notation left := left. -Ltac2 Notation "eleft" bnd(thunk(bindings)) := left0 true bnd. +Ltac2 Notation "eleft" bnd(thunk(with_bindings)) := left0 true bnd. Ltac2 Notation eleft := eleft. Ltac2 right0 ev bnd := enter_h ev Std.right bnd. -Ltac2 Notation "right" bnd(thunk(bindings)) := right0 false bnd. +Ltac2 Notation "right" bnd(thunk(with_bindings)) := right0 false bnd. Ltac2 Notation right := right. -Ltac2 Notation "eright" bnd(thunk(bindings)) := right0 true bnd. +Ltac2 Notation "eright" bnd(thunk(with_bindings)) := right0 true bnd. Ltac2 Notation eright := eright. Ltac2 constructor0 ev n bnd := @@ -168,22 +168,22 @@ Ltac2 constructor0 ev n bnd := Ltac2 Notation "constructor" := Control.enter (fun () => Std.constructor false). Ltac2 Notation constructor := constructor. -Ltac2 Notation "constructor" n(tactic) bnd(thunk(bindings)) := constructor0 false n bnd. +Ltac2 Notation "constructor" n(tactic) bnd(thunk(with_bindings)) := constructor0 false n bnd. Ltac2 Notation "econstructor" := Control.enter (fun () => Std.constructor true). Ltac2 Notation econstructor := econstructor. -Ltac2 Notation "econstructor" n(tactic) bnd(thunk(bindings)) := constructor0 true n bnd. +Ltac2 Notation "econstructor" n(tactic) bnd(thunk(with_bindings)) := constructor0 true n bnd. Ltac2 elim0 ev c bnd use := let f ev ((c, bnd, use)) := Std.elim ev (c, bnd) use in enter_h ev f (fun () => c (), bnd (), use ()). -Ltac2 Notation "elim" c(thunk(constr)) bnd(thunk(bindings)) - use(thunk(opt(seq("using", constr, bindings)))) := +Ltac2 Notation "elim" c(thunk(constr)) bnd(thunk(with_bindings)) + use(thunk(opt(seq("using", constr, with_bindings)))) := elim0 false c bnd use. -Ltac2 Notation "eelim" c(thunk(constr)) bnd(thunk(bindings)) - use(thunk(opt(seq("using", constr, bindings)))) := +Ltac2 Notation "eelim" c(thunk(constr)) bnd(thunk(with_bindings)) + use(thunk(opt(seq("using", constr, with_bindings)))) := elim0 true c bnd use. Ltac2 apply0 adv ev cb cl := @@ -202,12 +202,12 @@ Ltac2 apply0 adv ev cb cl := Std.apply adv ev cb cl. Ltac2 Notation "eapply" - cb(list1(thunk(seq(constr, bindings)), ",")) + cb(list1(thunk(seq(constr, with_bindings)), ",")) cl(opt(seq(keyword("in"), ident, opt(seq(keyword("as"), intropattern))))) := apply0 true true cb cl. Ltac2 Notation "apply" - cb(list1(thunk(seq(constr, bindings)), ",")) + cb(list1(thunk(seq(constr, with_bindings)), ",")) cl(opt(seq(keyword("in"), ident, opt(seq(keyword("as"), intropattern))))) := apply0 true false cb cl. @@ -217,12 +217,12 @@ Ltac2 induction0 ev ic use := Ltac2 Notation "induction" ic(list1(induction_clause, ",")) - use(thunk(opt(seq("using", constr, bindings)))) := + use(thunk(opt(seq("using", constr, with_bindings)))) := induction0 false ic use. Ltac2 Notation "einduction" ic(list1(induction_clause, ",")) - use(thunk(opt(seq("using", constr, bindings)))) := + use(thunk(opt(seq("using", constr, with_bindings)))) := induction0 true ic use. Ltac2 destruct0 ev ic use := @@ -231,12 +231,12 @@ Ltac2 destruct0 ev ic use := Ltac2 Notation "destruct" ic(list1(induction_clause, ",")) - use(thunk(opt(seq("using", constr, bindings)))) := + use(thunk(opt(seq("using", constr, with_bindings)))) := destruct0 false ic use. Ltac2 Notation "edestruct" ic(list1(induction_clause, ",")) - use(thunk(opt(seq("using", constr, bindings)))) := + use(thunk(opt(seq("using", constr, with_bindings)))) := destruct0 true ic use. Ltac2 default_on_concl cl := -- cgit v1.2.3 From 8e6338d862873d7e377f59664bbd89e16c0a7309 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 25 Aug 2017 18:03:54 +0200 Subject: Introducing a distinct bindings scope. --- src/g_ltac2.ml4 | 7 +++++-- src/tac2core.ml | 1 + src/tac2entries.ml | 1 + src/tac2entries.mli | 1 + 4 files changed, 8 insertions(+), 2 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 95fcf79000..e7ab574747 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -327,9 +327,9 @@ open Tac2entries.Pltac let loc_of_ne_list l = Loc.merge_opt (fst (List.hd l)) (fst (List.last l)) GEXTEND Gram - GLOBAL: q_ident q_with_bindings q_intropattern q_intropatterns q_induction_clause + GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause q_rewriting q_clause q_dispatch q_occurrences q_strategy_flag - q_reference; + q_reference q_with_bindings; anti: [ [ "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] ; @@ -365,6 +365,9 @@ GEXTEND Gram Loc.tag ~loc:!@loc @@ QImplicitBindings bl ] ] ; + q_bindings: + [ [ bl = bindings -> bl ] ] + ; q_with_bindings: [ [ bl = with_bindings -> bl ] ] ; diff --git a/src/tac2core.ml b/src/tac2core.ml index 342e8f51e8..b67d70a5cb 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -912,6 +912,7 @@ let add_expr_scope name entry f = end let () = add_expr_scope "ident" q_ident (fun id -> Tac2quote.of_anti Tac2quote.of_ident id) +let () = add_expr_scope "bindings" q_bindings Tac2quote.of_bindings let () = add_expr_scope "with_bindings" q_with_bindings Tac2quote.of_bindings let () = add_expr_scope "intropattern" q_intropattern Tac2quote.of_intro_pattern let () = add_expr_scope "intropatterns" q_intropatterns Tac2quote.of_intro_patterns diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 6c8c2348fe..91c8d10e2d 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -25,6 +25,7 @@ struct let tac2expr = Pcoq.Gram.entry_create "tactic:tac2expr" let q_ident = Pcoq.Gram.entry_create "tactic:q_ident" +let q_bindings = Pcoq.Gram.entry_create "tactic:q_bindings" let q_with_bindings = Pcoq.Gram.entry_create "tactic:q_with_bindings" let q_intropattern = Pcoq.Gram.entry_create "tactic:q_intropattern" let q_intropatterns = Pcoq.Gram.entry_create "tactic:q_intropatterns" diff --git a/src/tac2entries.mli b/src/tac2entries.mli index d22ae7a953..2c5862b149 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -60,6 +60,7 @@ val tac2expr : raw_tacexpr Pcoq.Gram.entry open Tac2qexpr val q_ident : Id.t located or_anti Pcoq.Gram.entry +val q_bindings : bindings Pcoq.Gram.entry val q_with_bindings : bindings Pcoq.Gram.entry val q_intropattern : intro_pattern Pcoq.Gram.entry val q_intropatterns : intro_pattern list located Pcoq.Gram.entry -- cgit v1.2.3 From b3471b2bf449041b47c19e8e12249e4bb36af3ec Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 25 Aug 2017 18:06:36 +0200 Subject: Adding more notations for the lulz. --- tests/example2.v | 16 ++++++++++++++++ theories/Notations.v | 21 +++++++++++++++++---- 2 files changed, 33 insertions(+), 4 deletions(-) diff --git a/tests/example2.v b/tests/example2.v index 95485305dc..a7eb02050b 100644 --- a/tests/example2.v +++ b/tests/example2.v @@ -192,3 +192,19 @@ unfold &x at 1. let x := reference:(Nat.sub) in unfold Nat.add, $x in x. reflexivity. Qed. + +Goal exists x y : nat, x = y. +Proof. +exists 0, 0; reflexivity. +Qed. + +Goal exists x y : nat, x = y. +Proof. +eexists _, 0; reflexivity. +Qed. + +Goal exists x y : nat, x = y. +Proof. +refine '(let x := 0 in _). +eexists; exists &x; reflexivity. +Qed. diff --git a/theories/Notations.v b/theories/Notations.v index 8345344d94..2a496b3faf 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -140,12 +140,11 @@ Ltac2 exists0 ev bnds := match bnds with aux bnds end. -(* -Ltac2 Notation "exists" bnd(list0(thunk(with_bindings), ",")) := exists0 false bnd. +Ltac2 Notation "exists" bnd(list0(thunk(bindings), ",")) := exists0 false bnd. +(* Ltac2 Notation exists := exists. *) -Ltac2 Notation "eexists" bnd(list0(thunk(with_bindings), ",")) := exists0 true bnd. +Ltac2 Notation "eexists" bnd(list0(thunk(bindings), ",")) := exists0 true bnd. Ltac2 Notation eexists := eexists. -*) Ltac2 left0 ev bnd := enter_h ev Std.left bnd. @@ -308,6 +307,20 @@ Ltac2 Notation "erewrite" (** coretactics *) +Ltac2 exact0 ev c := + Control.enter (fun _ => + match ev with + | true => + let c := c () in + Control.refine (fun _ => c) + | false => + Control.with_holes c (fun c => Control.refine (fun _ => c)) + end + ). + +Ltac2 Notation "exact" c(thunk(open_constr)) := exact0 false c. +Ltac2 Notation "eexact" c(thunk(open_constr)) := exact0 true c. + Ltac2 Notation reflexivity := Std.reflexivity (). Ltac2 Notation assumption := Std.assumption (). -- cgit v1.2.3 From 126dc656963a7feb589b2a3574f0c55ad84d5f69 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 26 Aug 2017 00:52:39 +0200 Subject: Allowing to insert calls to Ltac1 references in Ltac2. --- doc/ltac2.md | 11 +++++++++++ src/g_ltac2.ml4 | 5 +++++ src/tac2core.ml | 14 ++++++++++++++ 3 files changed, 30 insertions(+) diff --git a/doc/ltac2.md b/doc/ltac2.md index a645331e2d..51e3ab664d 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -692,6 +692,17 @@ foo 0 ↦ (fun x => x ()) (fun _ => 0) Note that abbreviations are not typechecked at all, and may result in typing errors after expansion. +# Compatibility layer with Ltac1 + +## Ltac1 from Ltac2 + +One can call Ltac1 code from Ltac2 by using the `ltac1` quotation. It parses +a Ltac1 expression, and semantics of this quotation is the evaluation of the +corresponding code for its side effects. + +Beware, Ltac1 **cannot** access variables from the Ltac2 scope. One is limited +to the use of standalone function calls. + # Transition from Ltac1 Owing to the use of a bunch of notations, the transition shouldn't be diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index e7ab574747..6822b8e7ba 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -72,10 +72,14 @@ let tac2def_ext = Gram.entry_create "tactic:tac2def_ext" let tac2def_syn = Gram.entry_create "tactic:tac2def_syn" let tac2mode = Gram.entry_create "vernac:ltac2_command" +(** FUCK YOU API *) +let ltac1_expr = (Obj.magic Pltac.tactic_expr : Tacexpr.raw_tactic_expr Gram.entry) + let inj_wit wit loc x = CTacExt (loc, Genarg.in_gen (Genarg.rawwit wit) x) let inj_open_constr loc c = inj_wit Stdarg.wit_open_constr loc c let inj_pattern loc c = inj_wit Tac2env.wit_pattern loc c let inj_reference loc c = inj_wit Tac2env.wit_reference loc c +let inj_ltac1 loc e = inj_wit Tacarg.wit_tactic loc e let pattern_of_qualid loc id = if Tac2env.is_constructor (snd id) then CPatRef (loc, RelId id, []) @@ -170,6 +174,7 @@ GEXTEND Gram | IDENT "ident"; ":"; "("; c = lident; ")" -> Tac2quote.of_ident c | IDENT "pattern"; ":"; "("; c = Constr.lconstr_pattern; ")" -> inj_pattern !@loc c | IDENT "reference"; ":"; "("; c = globref; ")" -> inj_reference !@loc c + | IDENT "ltac1"; ":"; "("; qid = ltac1_expr; ")" -> inj_ltac1 !@loc qid ] ] ; let_clause: diff --git a/src/tac2core.ml b/src/tac2core.ml index b67d70a5cb..b95410f40e 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -737,6 +737,20 @@ let () = } in define_ml_object Tac2env.wit_reference obj +let () = + let interp _ tac = + (** FUCK YOU API *) + (Obj.magic Ltac_plugin.Tacinterp.eval_tactic tac : unit Proofview.tactic) >>= fun () -> + return v_unit + in + let obj = { + ml_type = t_unit; + ml_interp = interp; + } in + define_ml_object Ltac_plugin.Tacarg.wit_tactic obj + +(** Ltac2 in terms *) + let () = let interp ist env sigma concl tac = let fold id (Val.Dyn (tag, v)) (accu : environment) : environment = -- cgit v1.2.3 From bec2a0ad6eb60d33b5e3ab613d108f456df42a49 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Sat, 26 Aug 2017 09:40:31 +0200 Subject: Typos --- doc/ltac2.md | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/doc/ltac2.md b/doc/ltac2.md index 51e3ab664d..5b1776b64f 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -488,7 +488,7 @@ Ltac2 Definition myconstr () := constr:(nat -> 0). ``` Term antiquotations are type-checked in the enclosing Ltac2 typing context -of the corresponding term expression. For instance, the following with +of the corresponding term expression. For instance, the following will type-check. ``` @@ -523,7 +523,7 @@ This pattern is so common that we provide dedicated Ltac2 and Coq term notations for it. - `&x` as an Ltac2 expression expands to `hyp @x`. -- `&x` as an Coq constr expression expands to +- `&x` as a Coq constr expression expands to `ltac2:(refine (fun () => hyp @x))`. #### Dynamic semantics @@ -538,7 +538,7 @@ under focus, with the hypotheses coming from the current environment extended with the bound variables of the term, and the resulting term is fed into the quoted term. -Relative orders of evaluation of antiquotations and quoted term is not +Relative orders of evaluation of antiquotations and quoted term are not specified. For instance, in the following example, `tac` will be evaluated in a context -- cgit v1.2.3 From 7f562a9539522e56004596a751758a08cee798b1 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 26 Aug 2017 16:58:06 +0200 Subject: Allowing calls to Ltac2 inside Ltac1. --- doc/ltac2.md | 22 ++++++++++++++++++++-- src/tac2core.ml | 20 ++++++++++++++++++++ tests/compat.v | 56 ++++++++++++++++++++++++++++++++++++++++++++++++++++++++ 3 files changed, 96 insertions(+), 2 deletions(-) create mode 100644 tests/compat.v diff --git a/doc/ltac2.md b/doc/ltac2.md index 5b1776b64f..d1c5c68494 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -698,11 +698,30 @@ errors after expansion. One can call Ltac1 code from Ltac2 by using the `ltac1` quotation. It parses a Ltac1 expression, and semantics of this quotation is the evaluation of the -corresponding code for its side effects. +corresponding code for its side effects. In particular, in cannot return values, +and the quotation has type `unit`. Beware, Ltac1 **cannot** access variables from the Ltac2 scope. One is limited to the use of standalone function calls. +## Ltac2 from Ltac1 + +Same as above by switching Ltac1 by Ltac2 and using the `ltac2` quotation +instead. + +Note that the tactic expression is evaluated eagerly, if one wants to use it as +an argument to a Ltac1 function, she has to resort to the good old +`idtac; ltac2:(foo)` trick. For instance, the code below will fail immediately +and won't print anything. + +``` +Ltac mytac tac := idtac "wow"; tac. + +Goal True. +Proof. +mytac ltac2:(fail). +``` + # Transition from Ltac1 Owing to the use of a bunch of notations, the transition shouldn't be @@ -812,6 +831,5 @@ your duty to catch it and reraise it depending on your use. # TODO - Implement deep pattern-matching. -- Implement compatibility layer with Ltac1 - Craft an expressive set of primitive functions - Implement native compilation to OCaml diff --git a/src/tac2core.ml b/src/tac2core.ml index b95410f40e..118bea0f8e 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -765,6 +765,26 @@ let () = in Pretyping.register_constr_interp0 wit_ltac2 interp +(** Ltac2 in Ltac1 *) + +let () = + (** FUCK YOU API *) + let e = (Obj.magic Tac2entries.Pltac.tac2expr : _ API.Pcoq.Gram.entry) in + let inject (loc, v) = Tacexpr.TacGeneric (in_gen (rawwit wit_ltac2) v) in + Ltac_plugin.Tacentries.create_ltac_quotation "ltac2" inject (e, None) + +let () = + let open Ltac_plugin in + let open Tacinterp in + let idtac = Value.of_closure (default_ist ()) (Tacexpr.TacId []) in + (** FUCK YOU API *) + let idtac = (Obj.magic idtac : Geninterp.Val.t) in + let interp ist tac = + Tac2interp.interp Tac2interp.empty_environment tac >>= fun _ -> + Ftactic.return idtac + in + Geninterp.register_interp0 wit_ltac2 interp + (** Patterns *) let () = diff --git a/tests/compat.v b/tests/compat.v new file mode 100644 index 0000000000..44421349da --- /dev/null +++ b/tests/compat.v @@ -0,0 +1,56 @@ +Require Import Ltac2.Ltac2. + +Import Ltac2.Notations. + +(** Test calls to Ltac1 from Ltac2 *) + +Ltac2 foo () := ltac1:(discriminate). + +Goal true = false -> False. +Proof. +foo (). +Qed. + +Goal true = false -> false = true. +Proof. +intros H; ltac1:(match goal with [ H : ?P |- _ ] => rewrite H end); reflexivity. +Qed. + +Goal true = false -> false = true. +Proof. +(** FIXME when the non-strict mode is implemented. *) +Fail intros H; ltac1:(rewrite H); reflexivity. +Abort. + +(** Variables do not cross the compatibility layer boundary. *) +Fail Ltac2 bar nay := ltac1:(discriminate nay). + +(** Test calls to Ltac2 from Ltac1 *) + +Set Default Proof Mode "Classic". + +Ltac foo := ltac2:(foo ()). + +Goal true = false -> False. +Proof. +ltac2:(foo ()). +Qed. + +Goal true = false -> False. +Proof. +foo. +Qed. + +(** Variables do not cross the compatibility layer boundary. *) +Fail Ltac bar x := ltac2:(foo x). + +Ltac mytac tac := idtac "wow". + +Goal True. +Proof. +(** Fails because quotation is evaluated eagerly *) +Fail mytac ltac2:(fail). +(** One has to thunk thanks to the idtac trick *) +let t := idtac; ltac2:(fail) in mytac t. +constructor. +Qed. -- cgit v1.2.3 From 9bbdee3c09c92654bb8937b9939a9b9c69c23d1d Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 27 Aug 2017 01:04:19 +0200 Subject: Introducing rebindable toplevel definitions. --- doc/ltac2.md | 13 ++++++++- src/g_ltac2.ml4 | 20 +++++++++---- src/tac2entries.ml | 84 +++++++++++++++++++++++++++++++++++++++++++++++++---- src/tac2entries.mli | 2 +- src/tac2env.ml | 14 ++++++--- src/tac2env.mli | 10 +++++-- src/tac2expr.mli | 5 ++-- src/tac2intern.ml | 29 ++++++++++++++---- src/tac2intern.mli | 4 +++ src/tac2interp.ml | 2 +- tests/rebind.v | 24 +++++++++++++++ 11 files changed, 180 insertions(+), 27 deletions(-) create mode 100644 tests/rebind.v diff --git a/doc/ltac2.md b/doc/ltac2.md index d1c5c68494..5c057b3ead 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -218,7 +218,7 @@ Limitations: for now, deep pattern matching is not implemented yet. One can define a new global Ltac2 value using the following syntax. ``` VERNAC ::= -| "Ltac2" RECFLAG LIDENT ":=" TERM +| "Ltac2" MUTFLAG RECFLAG LIDENT ":=" TERM ``` For semantic reasons, the body of the Ltac2 definition must be a syntactical @@ -227,6 +227,17 @@ values. If the `RECFLAG` is set, the tactic is expanded into a recursive binding. +If the `MUTFLAG` is set, the definition can be redefined at a later stage. This +can be performed through the following command. + +``` +VERNAC ::= +| "Ltac2" "Set" QUALID ":=" TERM +``` + +Mutable definitions act like dynamic binding, i.e. at runtime, the last defined +value for this entry is chosen. This is useful for global flags and the like. + ## Reduction We use the usual ML call-by-value reduction, with an otherwise unspecified diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 6822b8e7ba..ae7e255896 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -70,6 +70,7 @@ let tac2def_val = Gram.entry_create "tactic:tac2def_val" let tac2def_typ = Gram.entry_create "tactic:tac2def_typ" let tac2def_ext = Gram.entry_create "tactic:tac2def_ext" let tac2def_syn = Gram.entry_create "tactic:tac2def_syn" +let tac2def_mut = Gram.entry_create "tactic:tac2def_mut" let tac2mode = Gram.entry_create "vernac:ltac2_command" (** FUCK YOU API *) @@ -90,7 +91,8 @@ let pattern_of_qualid loc id = CErrors.user_err ~loc (Pp.str "Syntax error") GEXTEND Gram - GLOBAL: tac2expr tac2type tac2def_val tac2def_typ tac2def_ext tac2def_syn; + GLOBAL: tac2expr tac2type tac2def_val tac2def_typ tac2def_ext tac2def_syn + tac2def_mut; tac2pat: [ "1" LEFTA [ id = Prim.qualid; pl = LIST1 tac2pat LEVEL "0" -> @@ -158,6 +160,10 @@ GEXTEND Gram [ [ IDENT "rec" -> true | -> false ] ] ; + mut_flag: + [ [ IDENT "mutable" -> true + | -> false ] ] + ; typ_param: [ [ "'"; id = Prim.ident -> id ] ] ; @@ -228,10 +234,13 @@ GEXTEND Gram ] ] ; tac2def_val: - [ [ isrec = rec_flag; l = LIST1 tac2def_body SEP "with" -> - StrVal (isrec, l) + [ [ mut = mut_flag; isrec = rec_flag; l = LIST1 tac2def_body SEP "with" -> + StrVal (mut, isrec, l) ] ] ; + tac2def_mut: + [ [ "Set"; qid = Prim.qualid; ":="; e = tac2expr -> StrMut (qid, e) ] ] + ; tac2typ_knd: [ [ t = tac2type -> CTydDef (Some t) | "["; ".."; "]" -> CTydOpn @@ -253,7 +262,7 @@ GEXTEND Gram | -> [] ] ] ; tac2rec_field: - [ [ mut = [ -> false | IDENT "mutable" -> true]; id = Prim.ident; ":"; t = tac2type -> (id, mut, t) ] ] + [ [ mut = mut_flag; id = Prim.ident; ":"; t = tac2type -> (id, mut, t) ] ] ; tac2rec_fieldexprs: [ [ f = tac2rec_fieldexpr; ";"; l = tac2rec_fieldexprs -> f :: l @@ -653,11 +662,12 @@ PRINTED BY pr_ltac2entry | [ tac2def_typ(t) ] -> [ t ] | [ tac2def_ext(e) ] -> [ e ] | [ tac2def_syn(e) ] -> [ e ] +| [ tac2def_mut(e) ] -> [ e ] END let classify_ltac2 = function | StrSyn _ -> Vernacexpr.VtUnknown, Vernacexpr.VtNow -| StrVal _ | StrPrm _ | StrTyp _ -> Vernac_classifier.classify_as_sideeff +| StrMut _ | StrVal _ | StrPrm _ | StrTyp _ -> Vernac_classifier.classify_as_sideeff VERNAC COMMAND EXTEND VernacDeclareTactic2Definition | [ "Ltac2" ltac2_entry(e) ] => [ classify_ltac2 e ] -> [ diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 91c8d10e2d..da7c07c134 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -42,20 +42,31 @@ end type tacdef = { tacdef_local : bool; + tacdef_mutable : bool; tacdef_expr : glb_tacexpr; tacdef_type : type_scheme; } let perform_tacdef visibility ((sp, kn), def) = let () = if not def.tacdef_local then Tac2env.push_ltac visibility sp (TacConstant kn) in - Tac2env.define_global kn (def.tacdef_expr, def.tacdef_type) + let data = { + Tac2env.gdata_expr = def.tacdef_expr; + gdata_type = def.tacdef_type; + gdata_mutable = def.tacdef_mutable; + } in + Tac2env.define_global kn data let load_tacdef i obj = perform_tacdef (Until i) obj let open_tacdef i obj = perform_tacdef (Exactly i) obj let cache_tacdef ((sp, kn), def) = let () = Tac2env.push_ltac (Until 1) sp (TacConstant kn) in - Tac2env.define_global kn (def.tacdef_expr, def.tacdef_type) + let data = { + Tac2env.gdata_expr = def.tacdef_expr; + gdata_type = def.tacdef_type; + gdata_mutable = def.tacdef_mutable; + } in + Tac2env.define_global kn data let subst_tacdef (subst, def) = let expr' = subst_expr subst def.tacdef_expr in @@ -296,7 +307,7 @@ let inline_rec_tactic tactics = in List.map map tactics -let register_ltac ?(local = false) isrec tactics = +let register_ltac ?(local = false) ?(mut = false) isrec tactics = let map ((loc, na), e) = let id = match na with | Anonymous -> @@ -329,6 +340,7 @@ let register_ltac ?(local = false) isrec tactics = let iter (id, e, t) = let def = { tacdef_local = local; + tacdef_mutable = mut; tacdef_expr = e; tacdef_type = t; } in @@ -423,6 +435,7 @@ let register_primitive ?(local = false) (loc, id) t ml = let e = GTacFun (bnd, GTacPrm (ml, arg)) in let def = { tacdef_local = local; + tacdef_mutable = false; tacdef_expr = e; tacdef_type = t; } in @@ -659,13 +672,72 @@ let register_notation ?(local = false) tkn lev body = match tkn, lev with } in Lib.add_anonymous_leaf (inTac2Notation ext) +type redefinition = { + redef_kn : ltac_constant; + redef_body : glb_tacexpr; +} + +let perform_redefinition (_, redef) = + let kn = redef.redef_kn in + let data, _ = Tac2env.interp_global kn in + let data = { data with Tac2env.gdata_expr = redef.redef_body } in + Tac2env.define_global kn data + +let subst_redefinition (subst, redef) = + let kn = Mod_subst.subst_kn subst redef.redef_kn in + let body = Tac2intern.subst_expr subst redef.redef_body in + if kn == redef.redef_kn && body == redef.redef_body then redef + else { redef_kn = kn; redef_body = body } + +let classify_redefinition o = Substitute o + +let inTac2Redefinition : redefinition -> obj = + declare_object {(default_object "TAC2-REDEFINITION") with + cache_function = perform_redefinition; + open_function = (fun _ -> perform_redefinition); + subst_function = subst_redefinition; + classify_function = classify_redefinition } + +let register_redefinition ?(local = false) (loc, qid) e = + let kn = + try Tac2env.locate_ltac qid + with Not_found -> user_err ?loc (str "Unknown tactic " ++ pr_qualid qid) + in + let kn = match kn with + | TacConstant kn -> kn + | TacAlias _ -> + user_err ?loc (str "Cannot redefine syntactic abbreviations") + in + let (data, _) = Tac2env.interp_global kn in + let () = + if not (data.Tac2env.gdata_mutable) then + user_err ?loc (str "The tactic " ++ pr_qualid qid ++ str " is not declared as mutable") + in + let (e, t) = intern e in + let () = + if not (is_value e) then + user_err ?loc (str "Tactic definition must be a syntactical value") + in + let () = + if not (Tac2intern.check_subtype t data.Tac2env.gdata_type) then + let name = int_name () in + user_err ?loc (str "Type " ++ pr_glbtype name (snd t) ++ + str " is not a subtype of " ++ pr_glbtype name (snd data.Tac2env.gdata_type)) + in + let def = { + redef_kn = kn; + redef_body = e; + } in + Lib.add_anonymous_leaf (inTac2Redefinition def) + (** Toplevel entries *) let register_struct ?local str = match str with -| StrVal (isrec, e) -> register_ltac ?local isrec e +| StrVal (mut, isrec, e) -> register_ltac ?local ~mut isrec e | StrTyp (isrec, t) -> register_type ?local isrec t | StrPrm (id, t, ml) -> register_primitive ?local id t ml | StrSyn (tok, lev, e) -> register_notation ?local tok lev e +| StrMut (qid, e) -> register_redefinition ?local qid e (** Printing *) @@ -685,7 +757,9 @@ let print_ltac ref = in match kn with | TacConstant kn -> - let (e, _, (_, t)) = Tac2env.interp_global kn in + let data, _ = Tac2env.interp_global kn in + let e = data.Tac2env.gdata_expr in + let (_, t) = data.Tac2env.gdata_type in let name = int_name () in Feedback.msg_notice ( hov 0 ( diff --git a/src/tac2entries.mli b/src/tac2entries.mli index 2c5862b149..acb99a34b1 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -13,7 +13,7 @@ open Tac2expr (** {5 Toplevel definitions} *) -val register_ltac : ?local:bool -> rec_flag -> +val register_ltac : ?local:bool -> ?mut:bool -> rec_flag -> (Name.t located * raw_tacexpr) list -> unit val register_type : ?local:bool -> rec_flag -> diff --git a/src/tac2env.ml b/src/tac2env.ml index 65276ec57f..59344e336b 100644 --- a/src/tac2env.ml +++ b/src/tac2env.ml @@ -12,6 +12,12 @@ open Names open Libnames open Tac2expr +type global_data = { + gdata_expr : glb_tacexpr; + gdata_type : type_scheme; + gdata_mutable : bool; +} + type constructor_data = { cdata_prms : int; cdata_type : type_constant; @@ -28,7 +34,7 @@ type projection_data = { } type ltac_state = { - ltac_tactics : (glb_tacexpr * type_scheme) KNmap.t; + ltac_tactics : global_data KNmap.t; ltac_constructors : constructor_data KNmap.t; ltac_projections : projection_data KNmap.t; ltac_types : glb_quant_typedef KNmap.t; @@ -49,7 +55,7 @@ let ltac_state = Summary.ref empty_state ~name:"ltac2-state" let rec eval_pure = function | GTacAtm (AtmInt n) -> ValInt n | GTacRef kn -> - let (e, _) = + let { gdata_expr = e } = try KNmap.find kn ltac_state.contents.ltac_tactics with Not_found -> assert false in @@ -68,8 +74,8 @@ let define_global kn e = ltac_state := { state with ltac_tactics = KNmap.add kn e state.ltac_tactics } let interp_global kn = - let (e, t) = KNmap.find kn ltac_state.contents.ltac_tactics in - (e, eval_pure e, t) + let data = KNmap.find kn ltac_state.contents.ltac_tactics in + (data, eval_pure data.gdata_expr) let define_constructor kn t = let state = !ltac_state in diff --git a/src/tac2env.mli b/src/tac2env.mli index 20bf24d19d..8a5fb531d8 100644 --- a/src/tac2env.mli +++ b/src/tac2env.mli @@ -16,8 +16,14 @@ open Tac2expr (** {5 Toplevel definition of values} *) -val define_global : ltac_constant -> (glb_tacexpr * type_scheme) -> unit -val interp_global : ltac_constant -> (glb_tacexpr * valexpr * type_scheme) +type global_data = { + gdata_expr : glb_tacexpr; + gdata_type : type_scheme; + gdata_mutable : bool; +} + +val define_global : ltac_constant -> global_data -> unit +val interp_global : ltac_constant -> global_data * valexpr (** {5 Toplevel definition of types} *) diff --git a/src/tac2expr.mli b/src/tac2expr.mli index 281ed6c81e..78611d51ca 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -154,7 +154,7 @@ type sexpr = (** {5 Toplevel statements} *) type strexpr = -| StrVal of rec_flag * (Name.t located * raw_tacexpr) list +| StrVal of mutable_flag * rec_flag * (Name.t located * raw_tacexpr) list (** Term definition *) | StrTyp of rec_flag * (qualid located * redef_flag * raw_quant_typedef) list (** Type definition *) @@ -162,7 +162,8 @@ type strexpr = (** External definition *) | StrSyn of sexpr list * int option * raw_tacexpr (** Syntactic extensions *) - +| StrMut of qualid located * raw_tacexpr + (** Redefinition of mutable globals *) (** {5 Dynamic semantics} *) diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 765be92103..ef0763ff8e 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -360,20 +360,20 @@ let eq_or_tuple eq t1 t2 = match t1, t2 with | Other o1, Other o2 -> eq o1 o2 | _ -> false -let rec unify env t1 t2 = match kind env t1, kind env t2 with +let rec unify0 env t1 t2 = match kind env t1, kind env t2 with | GTypVar id, t | t, GTypVar id -> unify_var env id t | GTypArrow (t1, u1), GTypArrow (t2, u2) -> - let () = unify env t1 t2 in - unify env u1 u2 + let () = unify0 env t1 t2 in + unify0 env u1 u2 | GTypRef (kn1, tl1), GTypRef (kn2, tl2) -> if eq_or_tuple KerName.equal kn1 kn2 then - List.iter2 (fun t1 t2 -> unify env t1 t2) tl1 tl2 + List.iter2 (fun t1 t2 -> unify0 env t1 t2) tl1 tl2 else raise (CannotUnify (t1, t2)) | _ -> raise (CannotUnify (t1, t2)) let unify ?loc env t1 t2 = - try unify env t1 t2 + try unify0 env t1 t2 with CannotUnify (u1, u2) -> user_err ?loc (str "This expression has type " ++ pr_glbtype env t1 ++ str " but an expression was expected of type " ++ pr_glbtype env t2) @@ -663,7 +663,7 @@ let rec intern_rec env = function let sch = Id.Map.find id env.env_var in (GTacVar id, fresh_mix_type_scheme env sch) | ArgArg (TacConstant kn) -> - let (_, _, sch) = Tac2env.interp_global kn in + let { Tac2env.gdata_type = sch }, _ = Tac2env.interp_global kn in (GTacRef kn, fresh_type_scheme env sch) | ArgArg (TacAlias kn) -> let e = Tac2env.interp_alias kn in @@ -1162,6 +1162,23 @@ let intern_open_type t = let t = normalize env (count, vars) t in (!count, t) +(** Subtyping *) + +let check_subtype t1 t2 = + let env = empty_env () in + let t1 = fresh_type_scheme env t1 in + (** We build a substitution mimicking rigid variable by using dummy refs *) + let mb = MBId.make DirPath.empty (Id.of_string "_t") in + let rigid i = + let kn = KerName.make (MPbound mb) DirPath.empty (Label.make "_t") in + GTypRef (Other kn, []) + in + let (n, t2) = t2 in + let subst = Array.init n rigid in + let substf i = subst.(i) in + let t2 = subst_type substf t2 in + try unify0 env t1 t2; true with CannotUnify _ -> false + (** Globalization *) let get_projection0 var = match var with diff --git a/src/tac2intern.mli b/src/tac2intern.mli index 898df649ba..dac074a0eb 100644 --- a/src/tac2intern.mli +++ b/src/tac2intern.mli @@ -22,6 +22,10 @@ val intern_open_type : raw_typexpr -> type_scheme val is_value : glb_tacexpr -> bool val check_unit : ?loc:Loc.t -> type_scheme -> unit +val check_subtype : type_scheme -> type_scheme -> bool +(** [check_subtype t1 t2] returns [true] iff all values of intances of type [t1] + also have type [t2]. *) + val subst_type : substitution -> 'a glb_typexpr -> 'a glb_typexpr val subst_expr : substitution -> glb_tacexpr -> glb_tacexpr val subst_quant_typedef : substitution -> glb_quant_typedef -> glb_quant_typedef diff --git a/src/tac2interp.ml b/src/tac2interp.ml index d3bc79957b..3e1a048d29 100644 --- a/src/tac2interp.ml +++ b/src/tac2interp.ml @@ -38,7 +38,7 @@ let get_var ist id = anomaly (str "Unbound variable " ++ Id.print id) let get_ref ist kn = - try pi2 (Tac2env.interp_global kn) with Not_found -> + try snd (Tac2env.interp_global kn) with Not_found -> anomaly (str "Unbound reference" ++ KerName.print kn) let return = Proofview.tclUNIT diff --git a/tests/rebind.v b/tests/rebind.v new file mode 100644 index 0000000000..270fdd0b69 --- /dev/null +++ b/tests/rebind.v @@ -0,0 +1,24 @@ +Require Import Ltac2.Ltac2 Ltac2.Notations. + +Ltac2 mutable foo () := constructor. + +Goal True. +Proof. +foo (). +Qed. + +Ltac2 Set foo := fun _ => fail. + +Goal True. +Proof. +Fail foo (). +constructor. +Qed. + +(** Not the right type *) +Fail Ltac2 Set foo := 0. + +Ltac2 bar () := (). + +(** Cannot redefine non-mutable tactics *) +Fail Ltac2 Set bar := fun _ => (). -- cgit v1.2.3 From c6d28beca01809dbd06b3b36ea53bd4a94824083 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 27 Aug 2017 17:07:15 +0200 Subject: Proper handling of rigid variables in subtyping. --- src/tac2intern.ml | 8 ++------ tests/rebind.v | 10 ++++++++++ 2 files changed, 12 insertions(+), 6 deletions(-) diff --git a/src/tac2intern.ml b/src/tac2intern.ml index ef0763ff8e..40e0ffb34e 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -1167,12 +1167,8 @@ let intern_open_type t = let check_subtype t1 t2 = let env = empty_env () in let t1 = fresh_type_scheme env t1 in - (** We build a substitution mimicking rigid variable by using dummy refs *) - let mb = MBId.make DirPath.empty (Id.of_string "_t") in - let rigid i = - let kn = KerName.make (MPbound mb) DirPath.empty (Label.make "_t") in - GTypRef (Other kn, []) - in + (** We build a substitution mimicking rigid variable by using dummy tuples *) + let rigid i = GTypRef (Tuple i, []) in let (n, t2) = t2 in let subst = Array.init n rigid in let substf i = subst.(i) in diff --git a/tests/rebind.v b/tests/rebind.v index 270fdd0b69..e1c20a2059 100644 --- a/tests/rebind.v +++ b/tests/rebind.v @@ -22,3 +22,13 @@ Ltac2 bar () := (). (** Cannot redefine non-mutable tactics *) Fail Ltac2 Set bar := fun _ => (). + +(** Subtype check *) + +Ltac2 mutable rec f x := f x. + +Fail Ltac2 Set f := fun x => x. + +Ltac2 mutable g x := x. + +Ltac2 Set g := f. -- cgit v1.2.3 From 4c822dbb1c01139e95c165515777703263806ec1 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 27 Aug 2017 17:17:00 +0200 Subject: Ensure no confusion with unit in rigid variables. --- src/tac2intern.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 40e0ffb34e..5c1c90e924 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -1168,7 +1168,7 @@ let check_subtype t1 t2 = let env = empty_env () in let t1 = fresh_type_scheme env t1 in (** We build a substitution mimicking rigid variable by using dummy tuples *) - let rigid i = GTypRef (Tuple i, []) in + let rigid i = GTypRef (Tuple (i + 1), []) in let (n, t2) = t2 in let subst = Array.init n rigid in let substf i = subst.(i) in -- cgit v1.2.3 From e430e9823960a136ee65c5977d89113574413449 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 27 Aug 2017 20:22:12 +0200 Subject: Fix semantics of the solve tactical. --- theories/Notations.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/Notations.v b/theories/Notations.v index 2a496b3faf..ea31d4ef47 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -84,7 +84,7 @@ Ltac2 rec solve0 tacs := match tacs with | [] => Control.zero Tactic_failure | tac :: tacs => - Control.enter (fun _ => orelse (fun _ => complete tac) (fun _ => first0 tacs)) + Control.enter (fun _ => orelse (fun _ => complete tac) (fun _ => solve0 tacs)) end. Ltac2 Notation "solve" "[" tacs(list0(thunk(tactic(6)), "|")) "]" := solve0 tacs. -- cgit v1.2.3 From 0a5097752646f5bf3fd542880d4e33ece771f588 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 27 Aug 2017 20:44:53 +0200 Subject: Notation for clear. --- theories/Notations.v | 8 +++++++- 1 file changed, 7 insertions(+), 1 deletion(-) diff --git a/theories/Notations.v b/theories/Notations.v index ea31d4ef47..46c0e5e79f 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -329,7 +329,13 @@ Ltac2 Notation etransitivity := Std.etransitivity (). Ltac2 Notation admit := Std.admit (). -Ltac2 Notation clear := Std.keep []. +Ltac2 clear0 ids := match ids with +| [] => Std.keep [] +| _ => Std.clear ids +end. + +Ltac2 Notation "clear" ids(list0(ident)) := clear0 ids. +Ltac2 Notation clear := clear. Ltac2 Notation refine := Control.refine. -- cgit v1.2.3 From 6bb9c33f0e35a694ca253bc766f9a235d2073a4f Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 27 Aug 2017 21:33:55 +0200 Subject: Do not reuse the Val.t type in toplevel values. --- src/tac2core.ml | 6 +++--- src/tac2expr.mli | 2 +- src/tac2ffi.ml | 6 +++--- 3 files changed, 7 insertions(+), 7 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 118bea0f8e..5b752840a4 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -691,7 +691,7 @@ let interp_constr flags ist (c, _) = Proofview.V82.wrap_exceptions begin fun () -> let ist = to_lvar ist in let (sigma, c) = understand_ltac flags env sigma ist WithoutTypeConstraint c in - let c = ValExt (Val.Dyn (Value.val_constr, c)) in + let c = ValExt (Value.val_constr, c) in Proofview.Unsafe.tclEVARS sigma >>= fun () -> Proofview.tclUNIT c end @@ -714,7 +714,7 @@ let () = define_ml_object Stdarg.wit_open_constr obj let () = - let interp _ id = return (ValExt (Val.Dyn (Value.val_ident, id))) in + let interp _ id = return (ValExt (Value.val_ident, id)) in let obj = { ml_type = t_ident; ml_interp = interp; @@ -722,7 +722,7 @@ let () = define_ml_object Stdarg.wit_ident obj let () = - let interp _ c = return (ValExt (Val.Dyn (Value.val_pattern, c))) in + let interp _ c = return (ValExt (Value.val_pattern, c)) in let obj = { ml_type = t_pattern; ml_interp = interp; diff --git a/src/tac2expr.mli b/src/tac2expr.mli index 78611d51ca..0c9112d902 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -186,7 +186,7 @@ type valexpr = (** Closures *) | ValOpn of KerName.t * valexpr array (** Open constructors *) -| ValExt of Geninterp.Val.t +| ValExt : 'a Geninterp.Val.typ * 'a -> valexpr (** Arbitrary data *) and closure = { diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index b506a578b1..a9a0f5a479 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -33,7 +33,7 @@ let val_univ = Val.create "ltac2:universe" let val_kont : (Exninfo.iexn -> valexpr Proofview.tactic) Val.typ = Val.create "ltac2:kont" -let extract_val (type a) (tag : a Val.typ) (Val.Dyn (tag', v)) : a = +let extract_val (type a) (type b) (tag : a Val.typ) (tag' : b Val.typ) (v : b) : a = match Val.eq tag tag' with | None -> assert false | Some Refl -> v @@ -78,10 +78,10 @@ let rec to_list f = function | _ -> assert false let of_ext tag c = - ValExt (Val.Dyn (tag, c)) + ValExt (tag, c) let to_ext tag = function -| ValExt e -> extract_val tag e +| ValExt (tag', e) -> extract_val tag tag' e | _ -> assert false let of_constr c = of_ext val_constr c -- cgit v1.2.3 From 99ec137899c2684da2a8c221f333e0e9adee2c48 Mon Sep 17 00:00:00 2001 From: gallais Date: Mon, 28 Aug 2017 14:33:17 +0200 Subject: typos --- doc/ltac2.md | 16 ++++++++-------- 1 file changed, 8 insertions(+), 8 deletions(-) diff --git a/doc/ltac2.md b/doc/ltac2.md index 5c057b3ead..dd0dc391c6 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -149,7 +149,7 @@ Records are product types with named fields and eliminated by projection. Likewise they can be recursive if the `RECFLAG` is set. Open variants are a special kind of variant types whose constructors are not -statically defined, but can instead by extended dynamically. A typical example +statically defined, but can instead be extended dynamically. A typical example is the standard `exn` type. Pattern-matching must always include a catch-all clause. They can be extended by the following command. @@ -358,21 +358,21 @@ If one sees thunks as lazy lists, then `zero` is the empty list and `plus` is list concatenation, while `case` is pattern-matching. The backtracking is first-class, i.e. one can write -`plus "x" (fun () => "y") : string` producing a backtracking string. +`plus (fun () => "x") (fun _ => "y") : string` producing a backtracking string. These operations are expected to satisfy a few equations, most notably that they form a monoid compatible with sequentialization. ``` plus t zero ≡ t () -plus (fun () -> zero e) f ≡ f e -plus (plus t f) g ≡ plus t (fun e -> plus (f e) g) +plus (fun () => zero e) f ≡ f e +plus (plus t f) g ≡ plus t (fun e => plus (f e) g) -case (fun () -> zero e) ≡ Err e -case (fun () -> plus (fun () -> t) f) ≡ Val t f +case (fun () => zero e) ≡ Err e +case (fun () => plus (fun () => t) f) ≡ Val t f let x := zero e in u ≡ zero e -let x := plus t f in u ≡ plus (fun () -> let x := t in u) (fun e -> let x := f e in u) +let x := plus t f in u ≡ plus (fun () => let x := t in u) (fun e => let x := f e in u) (t, u, f, g, e values) ``` @@ -809,7 +809,7 @@ Ltac1 relies on a crazy amount of dynamic trickery to be able to tell apart bound variables from terms, hypotheses and whatnot. There is no such thing in Ltac2, as variables are recognized statically and other constructions do not live in the same syntactic world. Due to the abuse of quotations, it can -sometimes be complicated to know what a mere identifier represent in a tactic +sometimes be complicated to know what a mere identifier represents in a tactic expression. We recommend tracking the context and letting the compiler spit typing errors to understand what is going on. -- cgit v1.2.3 From ee7ebf17245c65b630d5b7c01de8ad9253bd9523 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 28 Aug 2017 20:31:33 +0200 Subject: Fix coq/ltac2#6: Expected and actual types are flipped. --- src/tac2intern.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 5c1c90e924..9d22b18af4 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -947,7 +947,7 @@ and intern_case env loc e pl = in brT in - let () = unify ~loc:(loc_of_tacexpr br) env ret tbr in + let () = unify ~loc:(loc_of_tacexpr br) env tbr ret in intern_branch rem in let () = intern_branch pl in -- cgit v1.2.3 From ad730081c470705fed0c3741a0373090ab9c7d17 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 28 Aug 2017 22:00:44 +0200 Subject: Fix README. --- README.md | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/README.md b/README.md index 4d6879d8af..73785f6368 100644 --- a/README.md +++ b/README.md @@ -14,7 +14,7 @@ bug-ridden. Don't mistake this for a final product! Installation ============ -This should compile with Coq 8.7, assuming the `COQLIB` variable is set +This should compile with Coq master, assuming the `COQBIN` variable is set correctly. Standard procedures for `coq_makefile`-generated plugins apply. Demo -- cgit v1.2.3 From 8ba24b7342c3885145406e588859a3e1e356987d Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 28 Aug 2017 22:43:32 +0200 Subject: Binding primitives to generate fresh variables. --- _CoqProject | 1 + src/tac2core.ml | 45 ++++++++++++++++++++++++++++++++++++++++++++- theories/Fresh.v | 26 ++++++++++++++++++++++++++ 3 files changed, 71 insertions(+), 1 deletion(-) create mode 100644 theories/Fresh.v diff --git a/_CoqProject b/_CoqProject index 4116d17554..ffe1dda032 100644 --- a/_CoqProject +++ b/_CoqProject @@ -36,6 +36,7 @@ theories/Control.v theories/Message.v theories/Constr.v theories/Pattern.v +theories/Fresh.v theories/Std.v theories/Notations.v theories/Ltac2.v diff --git a/src/tac2core.ml b/src/tac2core.ml index 5b752840a4..4a4f9afd88 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -66,7 +66,8 @@ let of_rec_declaration (nas, ts, cs) = Value.of_array Value.of_constr ts, Value.of_array Value.of_constr cs) -let val_valexpr = Val.create "ltac2:valexpr" +let val_valexpr : valexpr Val.typ = Val.create "ltac2:valexpr" +let val_free : Id.Set.t Val.typ = Val.create "ltac2:free" (** Stdlib exceptions *) @@ -593,6 +594,43 @@ let prm_check_interrupt : ml_tactic = function | [_] -> Proofview.tclCHECKINTERRUPT >>= fun () -> return v_unit | _ -> assert false +(** Fresh *) + +let prm_free_union : ml_tactic = function +| [set1; set2] -> + let set1 = Value.to_ext val_free set1 in + let set2 = Value.to_ext val_free set2 in + let ans = Id.Set.union set1 set2 in + return (Value.of_ext val_free ans) +| _ -> assert false + +let prm_free_of_ids : ml_tactic = function +| [ids] -> + let ids = Value.to_list Value.to_ident ids in + let free = List.fold_right Id.Set.add ids Id.Set.empty in + return (Value.of_ext val_free free) +| _ -> assert false + +let prm_free_of_constr : ml_tactic = function +| [c] -> + let c = Value.to_constr c in + Proofview.tclEVARMAP >>= fun sigma -> + let rec fold accu c = match EConstr.kind sigma c with + | Constr.Var id -> Id.Set.add id accu + | _ -> EConstr.fold sigma fold accu c + in + let ans = fold Id.Set.empty c in + return (Value.of_ext val_free ans) +| _ -> assert false + +let prm_fresh : ml_tactic = function +| [avoid; id] -> + let avoid = Value.to_ext val_free avoid in + let id = Value.to_ident id in + let nid = Namegen.next_ident_away_from id (fun id -> Id.Set.mem id avoid) in + return (Value.of_ident nid) +| _ -> assert false + (** Registering *) let () = Tac2env.define_primitive (pname "print") prm_print @@ -656,6 +694,11 @@ let () = Tac2env.define_primitive (pname "abstract") prm_abstract let () = Tac2env.define_primitive (pname "time") prm_time let () = Tac2env.define_primitive (pname "check_interrupt") prm_check_interrupt +let () = Tac2env.define_primitive (pname "fresh_fresh") prm_fresh +let () = Tac2env.define_primitive (pname "fresh_free_union") prm_free_union +let () = Tac2env.define_primitive (pname "fresh_free_of_ids") prm_free_of_ids +let () = Tac2env.define_primitive (pname "fresh_free_of_constr") prm_free_of_constr + (** ML types *) let constr_flags () = diff --git a/theories/Fresh.v b/theories/Fresh.v new file mode 100644 index 0000000000..5e876bb077 --- /dev/null +++ b/theories/Fresh.v @@ -0,0 +1,26 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* t -> t := "ltac2" "fresh_free_union". + +Ltac2 @ external of_ids : ident list -> t := "ltac2" "fresh_free_of_ids". + +Ltac2 @ external of_constr : constr -> t := "ltac2" "fresh_free_of_constr". + +End Free. + +Ltac2 @ external fresh : Free.t -> ident -> ident := "ltac2" "fresh_fresh". +(** Generate a fresh identifier with the given base name which is not a + member of the provided set of free variables. *) -- cgit v1.2.3 From ece1cc059c26351d05a0ef41131c663c37cb7761 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 29 Aug 2017 00:40:40 +0200 Subject: Binding an unsafe substitution function. --- src/tac2core.ml | 10 ++++++++++ theories/Constr.v | 4 ++++ 2 files changed, 14 insertions(+) diff --git a/src/tac2core.ml b/src/tac2core.ml index 4a4f9afd88..0fe4bc5fde 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -366,6 +366,15 @@ let prm_constr_kind : ml_tactic = function end | _ -> assert false +let prm_constr_substnl : ml_tactic = function +| [subst; k; c] -> + let subst = Value.to_list Value.to_constr subst in + let k = Value.to_int k in + let c = Value.to_constr c in + let ans = EConstr.Vars.substnl subst k c in + return (Value.of_constr ans) +| _ -> assert false + (** Patterns *) let prm_pattern_matches : ml_tactic = function @@ -653,6 +662,7 @@ let () = Tac2env.define_primitive (pname "string_set") prm_string_set let () = Tac2env.define_primitive (pname "constr_type") prm_constr_type let () = Tac2env.define_primitive (pname "constr_equal") prm_constr_equal let () = Tac2env.define_primitive (pname "constr_kind") prm_constr_kind +let () = Tac2env.define_primitive (pname "constr_substnl") prm_constr_substnl let () = Tac2env.define_primitive (pname "pattern_matches") prm_pattern_matches let () = Tac2env.define_primitive (pname "pattern_matches_subterm") prm_pattern_matches_subterm diff --git a/theories/Constr.v b/theories/Constr.v index d7cd3b58a3..bb02d94531 100644 --- a/theories/Constr.v +++ b/theories/Constr.v @@ -40,4 +40,8 @@ Ltac2 Type kind := [ Ltac2 @ external kind : constr -> kind := "ltac2" "constr_kind". +Ltac2 @ external substnl : constr list -> int -> constr -> constr := "ltac2" "constr_substnl". +(** [substnl [r₁;...;rₙ] k c] substitutes in parallel [Rel(k+1); ...; Rel(k+n)] with + [r₁;...;rₙ] in [c]. *) + End Unsafe. -- cgit v1.2.3 From db03c10aaafd3c0128a5b7504f14d4b7aaca842e Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 29 Aug 2017 01:16:21 +0200 Subject: Implementing Ltac2 antiquotations in constr syntax. --- src/g_ltac2.ml4 | 14 ++++++++++++++ src/tac2quote.ml | 5 +++++ src/tac2quote.mli | 3 +++ tests/quot.v | 7 +++++++ 4 files changed, 29 insertions(+) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index ae7e255896..672db12f1d 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -64,6 +64,16 @@ let test_ampersand_ident = | _ -> err ()) | _ -> err ()) +let test_dollar_ident = + Gram.Entry.of_parser "test_dollar_ident" + (fun strm -> + match stream_nth 0 strm with + | IDENT "$" | KEYWORD "$" -> + (match stream_nth 1 strm with + | IDENT _ -> () + | _ -> err ()) + | _ -> err ()) + let tac2expr = Tac2entries.Pltac.tac2expr let tac2type = Gram.entry_create "tactic:tac2type" let tac2def_val = Gram.entry_create "tactic:tac2def_val" @@ -649,6 +659,10 @@ GEXTEND Gram let tac = Tac2quote.of_exact_hyp ~loc:!@loc (Loc.tag ~loc:!@loc id) in let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in CAst.make ~loc:!@loc (CHole (None, IntroAnonymous, Some arg)) + | test_dollar_ident; "$"; id = Prim.ident -> + let tac = Tac2quote.of_exact_var ~loc:!@loc (Loc.tag ~loc:!@loc id) in + let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in + CAst.make ~loc:!@loc (CHole (None, IntroAnonymous, Some arg)) ] ] ; END diff --git a/src/tac2quote.ml b/src/tac2quote.ml index dbd2fd0529..f87e370032 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -233,6 +233,11 @@ let of_exact_hyp ?loc id = let refine = CTacRef (AbsKn (TacConstant (control_core "refine"))) in CTacApp (loc, refine, [thunk (of_hyp ~loc id)]) +let of_exact_var ?loc id = + let loc = Option.default dummy_loc loc in + let refine = CTacRef (AbsKn (TacConstant (control_core "refine"))) in + CTacApp (loc, refine, [thunk (of_variable id)]) + let of_dispatch tacs = let loc = Option.default dummy_loc (fst tacs) in let default = function diff --git a/src/tac2quote.mli b/src/tac2quote.mli index 7f3d9dce6e..b2687f01a3 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -58,6 +58,9 @@ val of_hyp : ?loc:Loc.t -> Id.t located -> raw_tacexpr val of_exact_hyp : ?loc:Loc.t -> Id.t located -> raw_tacexpr (** id ↦ 'Control.refine (fun () => Control.hyp @id') *) +val of_exact_var : ?loc:Loc.t -> Id.t located -> raw_tacexpr +(** id ↦ 'Control.refine (fun () => Control.hyp @id') *) + val of_dispatch : dispatch -> raw_tacexpr val of_strategy_flag : strategy_flag -> raw_tacexpr diff --git a/tests/quot.v b/tests/quot.v index c9aa1f9d14..4fa9c4fa4e 100644 --- a/tests/quot.v +++ b/tests/quot.v @@ -7,3 +7,10 @@ Ltac2 ref1 () := reference:(nat). Ltac2 ref2 () := reference:(Datatypes.nat). Fail Ltac2 ref () := reference:(i_certainly_dont_exist). Fail Ltac2 ref () := reference:(And.Me.neither). + +Goal True. +Proof. +let x := constr:(I) in +let y := constr:((fun z => z) $x) in +Control.refine (fun _ => y). +Qed. -- cgit v1.2.3 From 9f79d601c0863d5144fc07c5cea0e03ef41d244b Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 29 Aug 2017 14:04:23 +0200 Subject: Factorizing code for declaration of primitive tactics. --- src/tac2core.ml | 382 ++++++++++++++++++++++---------------------------------- 1 file changed, 152 insertions(+), 230 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 0fe4bc5fde..609dd40587 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -53,6 +53,10 @@ open Core let v_unit = ValInt 0 +let to_block = function +| ValBlk (_, v) -> v +| _ -> assert false + let of_name c = match c with | Anonymous -> Value.of_option Value.of_ident None | Name id -> Value.of_option Value.of_ident (Some id) @@ -111,146 +115,159 @@ let pf_apply f = (** Primitives *) -(** Printing *) +let define0 name f = Tac2env.define_primitive (pname name) begin function +| [_] -> f +| _ -> assert false +end -let prm_print : ml_tactic = function -| [pp] -> wrap_unit (fun () -> Feedback.msg_notice (Value.to_pp pp)) +let define1 name f = Tac2env.define_primitive (pname name) begin function +| [x] -> f x | _ -> assert false +end -let prm_message_of_int : ml_tactic = function -| [ValInt s] -> return (Value.of_pp (int s)) +let define2 name f = Tac2env.define_primitive (pname name) begin function +| [x; y] -> f x y | _ -> assert false +end -let prm_message_of_string : ml_tactic = function -| [ValStr s] -> return (Value.of_pp (str (Bytes.to_string s))) +let define3 name f = Tac2env.define_primitive (pname name) begin function +| [x; y; z] -> f x y z | _ -> assert false +end -let prm_message_of_constr : ml_tactic = function -| [c] -> +(** Printing *) + +let () = define1 "print" begin fun pp -> + wrap_unit (fun () -> Feedback.msg_notice (Value.to_pp pp)) +end + +let () = define1 "message_of_int" begin fun n -> + let n = Value.to_int n in + return (Value.of_pp (int n)) +end + +let () = define1 "message_of_string" begin fun s -> + let s = Value.to_string s in + return (Value.of_pp (str (Bytes.to_string s))) +end + +let () = define1 "message_of_constr" begin fun c -> pf_apply begin fun env sigma -> let c = Value.to_constr c in let pp = Printer.pr_econstr_env env sigma c in return (Value.of_pp pp) end -| _ -> assert false +end -let prm_message_of_ident : ml_tactic = function -| [c] -> +let () = define1 "message_of_ident" begin fun c -> let c = Value.to_ident c in let pp = Id.print c in return (Value.of_pp pp) -| _ -> assert false +end -let prm_message_concat : ml_tactic = function -| [m1; m2] -> +let () = define2 "message_concat" begin fun m1 m2 -> let m1 = Value.to_pp m1 in let m2 = Value.to_pp m2 in return (Value.of_pp (Pp.app m1 m2)) -| _ -> assert false +end (** Array *) -let prm_array_make : ml_tactic = function -| [ValInt n; x] -> +let () = define2 "array_make" begin fun n x -> + let n = Value.to_int n in if n < 0 || n > Sys.max_array_length then throw err_outofbounds else wrap (fun () -> ValBlk (0, Array.make n x)) -| _ -> assert false +end -let prm_array_length : ml_tactic = function -| [ValBlk (_, v)] -> return (ValInt (Array.length v)) -| _ -> assert false +let () = define1 "array_length" begin fun v -> + let v = to_block v in + return (ValInt (Array.length v)) +end -let prm_array_set : ml_tactic = function -| [ValBlk (_, v); ValInt n; x] -> +let () = define3 "array_set" begin fun v n x -> + let v = to_block v in + let n = Value.to_int n in if n < 0 || n >= Array.length v then throw err_outofbounds else wrap_unit (fun () -> v.(n) <- x) -| _ -> assert false +end -let prm_array_get : ml_tactic = function -| [ValBlk (_, v); ValInt n] -> +let () = define2 "array_get" begin fun v n -> + let v = to_block v in + let n = Value.to_int n in if n < 0 || n >= Array.length v then throw err_outofbounds else wrap (fun () -> v.(n)) -| _ -> assert false +end (** Ident *) -let prm_ident_equal : ml_tactic = function -| [id1; id2] -> +let () = define2 "ident_equal" begin fun id1 id2 -> let id1 = Value.to_ident id1 in let id2 = Value.to_ident id2 in return (Value.of_bool (Id.equal id1 id2)) -| _ -> assert false +end -let prm_ident_to_string : ml_tactic = function -| [id] -> +let () = define1 "ident_to_string" begin fun id -> let id = Value.to_ident id in return (Value.of_string (Id.to_string id)) -| _ -> assert false +end -let prm_ident_of_string : ml_tactic = function -| [s] -> +let () = define1 "ident_of_string" begin fun s -> let s = Value.to_string s in let id = try Some (Id.of_string s) with _ -> None in return (Value.of_option Value.of_ident id) -| _ -> assert false +end (** Int *) -let prm_int_equal : ml_tactic = function -| [m; n] -> +let () = define2 "int_equal" begin fun m n -> return (Value.of_bool (Value.to_int m == Value.to_int n)) -| _ -> assert false +end -let binop f : ml_tactic = function -| [m; n] -> return (Value.of_int (f (Value.to_int m) (Value.to_int n))) -| _ -> assert false +let binop n f = define2 n begin fun m n -> + return (Value.of_int (f (Value.to_int m) (Value.to_int n))) +end -let prm_int_compare args = binop Int.compare args -let prm_int_add args = binop (+) args -let prm_int_sub args = binop (-) args -let prm_int_mul args = binop ( * ) args +let () = binop "int_compare" Int.compare +let () = binop "int_add" (+) +let () = binop "int_sub" (-) +let () = binop "int_mul" ( * ) -let prm_int_neg : ml_tactic = function -| [m] -> return (Value.of_int (~- (Value.to_int m))) -| _ -> assert false +let () = define1 "int_neg" begin fun m -> + return (Value.of_int (~- (Value.to_int m))) +end (** String *) -let prm_string_make : ml_tactic = function -| [n; c] -> +let () = define2 "string_make" begin fun n c -> let n = Value.to_int n in let c = Value.to_char c in if n < 0 || n > Sys.max_string_length then throw err_outofbounds else wrap (fun () -> Value.of_string (Bytes.make n c)) -| _ -> assert false +end -let prm_string_length : ml_tactic = function -| [s] -> +let () = define1 "string_length" begin fun s -> return (Value.of_int (Bytes.length (Value.to_string s))) -| _ -> assert false +end -let prm_string_set : ml_tactic = function -| [s; n; c] -> +let () = define3 "string_set" begin fun s n c -> let s = Value.to_string s in let n = Value.to_int n in let c = Value.to_char c in if n < 0 || n >= Bytes.length s then throw err_outofbounds else wrap_unit (fun () -> Bytes.set s n c) -| _ -> assert false +end -let prm_string_get : ml_tactic = function -| [s; n] -> +let () = define2 "string_get" begin fun s n -> let s = Value.to_string s in let n = Value.to_int n in if n < 0 || n >= Bytes.length s then throw err_outofbounds else wrap (fun () -> Value.of_char (Bytes.get s n)) -| _ -> assert false +end (** Terms *) (** constr -> constr *) -let prm_constr_type : ml_tactic = function -| [c] -> +let () = define1 "constr_type" begin fun c -> let c = Value.to_constr c in let get_type env sigma = Proofview.V82.wrap_exceptions begin fun () -> @@ -259,20 +276,18 @@ let prm_constr_type : ml_tactic = function Proofview.Unsafe.tclEVARS sigma <*> Proofview.tclUNIT t end in pf_apply get_type -| _ -> assert false +end (** constr -> constr *) -let prm_constr_equal : ml_tactic = function -| [c1; c2] -> +let () = define2 "constr_equal" begin fun c1 c2 -> let c1 = Value.to_constr c1 in let c2 = Value.to_constr c2 in Proofview.tclEVARMAP >>= fun sigma -> let b = EConstr.eq_constr sigma c1 c2 in Proofview.tclUNIT (Value.of_bool b) -| _ -> assert false +end -let prm_constr_kind : ml_tactic = function -| [c] -> +let () = define1 "constr_kind" begin fun c -> let open Constr in Proofview.tclEVARMAP >>= fun sigma -> let c = Value.to_constr c in @@ -364,21 +379,19 @@ let prm_constr_kind : ml_tactic = function Value.of_constr c; |]) end -| _ -> assert false +end -let prm_constr_substnl : ml_tactic = function -| [subst; k; c] -> +let () = define3 "constr_substnl" begin fun subst k c -> let subst = Value.to_list Value.to_constr subst in let k = Value.to_int k in let c = Value.to_constr c in let ans = EConstr.Vars.substnl subst k c in return (Value.of_constr ans) -| _ -> assert false +end (** Patterns *) -let prm_pattern_matches : ml_tactic = function -| [pat; c] -> +let () = define2 "pattern_matches" begin fun pat c -> let pat = Value.to_pattern pat in let c = Value.to_constr c in pf_apply begin fun env sigma -> @@ -394,10 +407,9 @@ let prm_pattern_matches : ml_tactic = function return (Value.of_list of_pair ans) end end -| _ -> assert false +end -let prm_pattern_matches_subterm : ml_tactic = function -| [pat; c] -> +let () = define2 "pattern_matches_subterm" begin fun pat c -> let pat = Value.to_pattern pat in let c = Value.to_constr c in let open Constr_matching in @@ -413,74 +425,66 @@ let prm_pattern_matches_subterm : ml_tactic = function let ans = Constr_matching.match_appsubterm env sigma pat c in of_ans ans end -| _ -> assert false +end -let prm_pattern_instantiate : ml_tactic = function -| [ctx; c] -> +let () = define2 "pattern_instantiate" begin fun ctx c -> let ctx = EConstr.Unsafe.to_constr (Value.to_constr ctx) in let c = EConstr.Unsafe.to_constr (Value.to_constr c) in let ans = Termops.subst_meta [Constr_matching.special_meta, c] ctx in return (Value.of_constr (EConstr.of_constr ans)) -| _ -> assert false +end (** Error *) -let prm_throw : ml_tactic = function -| [e] -> +let () = define1 "throw" begin fun e -> let (e, info) = Value.to_exn e in Proofview.tclLIFT (Proofview.NonLogical.raise ~info e) -| _ -> assert false +end (** Control *) (** exn -> 'a *) -let prm_zero : ml_tactic = function -| [e] -> +let () = define1 "zero" begin fun e -> let (e, info) = Value.to_exn e in Proofview.tclZERO ~info e -| _ -> assert false +end (** (unit -> 'a) -> (exn -> 'a) -> 'a *) -let prm_plus : ml_tactic = function -| [x; k] -> +let () = define2 "plus" begin fun x k -> Proofview.tclOR (thaw x) (fun e -> interp_app k [Value.of_exn e]) -| _ -> assert false +end (** (unit -> 'a) -> 'a *) -let prm_once : ml_tactic = function -| [f] -> Proofview.tclONCE (thaw f) -| _ -> assert false +let () = define1 "once" begin fun f -> + Proofview.tclONCE (thaw f) +end (** (unit -> unit) list -> unit *) -let prm_dispatch : ml_tactic = function -| [l] -> +let () = define1 "dispatch" begin fun l -> let l = Value.to_list (fun f -> Proofview.tclIGNORE (thaw f)) l in Proofview.tclDISPATCH l >>= fun () -> return v_unit -| _ -> assert false +end (** (unit -> unit) list -> (unit -> unit) -> (unit -> unit) list -> unit *) -let prm_extend : ml_tactic = function -| [lft; tac; rgt] -> +let () = define3 "extend" begin fun lft tac rgt -> let lft = Value.to_list (fun f -> Proofview.tclIGNORE (thaw f)) lft in let tac = Proofview.tclIGNORE (thaw tac) in let rgt = Value.to_list (fun f -> Proofview.tclIGNORE (thaw f)) rgt in Proofview.tclEXTEND lft tac rgt >>= fun () -> return v_unit -| _ -> assert false +end (** (unit -> unit) -> unit *) -let prm_enter : ml_tactic = function -| [f] -> +let () = define1 "enter" begin fun f -> let f = Proofview.tclIGNORE (thaw f) in Proofview.tclINDEPENDENT f >>= fun () -> return v_unit -| _ -> assert false +end let k_var = Id.of_string "k" let e_var = Id.of_string "e" let prm_apply_kont_h = pname "apply_kont" (** (unit -> 'a) -> ('a * ('exn -> 'a)) result *) -let prm_case : ml_tactic = function -| [f] -> +let () = define1 "case" begin fun f -> Proofview.tclCASE (thaw f) >>= begin function | Proofview.Next (x, k) -> let k = { @@ -491,52 +495,48 @@ let prm_case : ml_tactic = function return (ValBlk (0, [| Value.of_tuple [| x; ValCls k |] |])) | Proofview.Fail e -> return (ValBlk (1, [| Value.of_exn e |])) end -| _ -> assert false +end (** 'a kont -> exn -> 'a *) -let prm_apply_kont : ml_tactic = function -| [k; e] -> (Value.to_ext Value.val_kont k) (Value.to_exn e) -| _ -> assert false +let () = define2 "apply_kont" begin fun k e -> + (Value.to_ext Value.val_kont k) (Value.to_exn e) +end (** int -> int -> (unit -> 'a) -> 'a *) -let prm_focus : ml_tactic = function -| [i; j; tac] -> +let () = define3 "focus" begin fun i j tac -> let i = Value.to_int i in let j = Value.to_int j in Proofview.tclFOCUS i j (thaw tac) -| _ -> assert false +end (** unit -> unit *) -let prm_shelve : ml_tactic = function -| [_] -> Proofview.shelve >>= fun () -> return v_unit -| _ -> assert false +let () = define0 "shelve" begin + Proofview.shelve >>= fun () -> return v_unit +end (** unit -> unit *) -let prm_shelve_unifiable : ml_tactic = function -| [_] -> Proofview.shelve_unifiable >>= fun () -> return v_unit -| _ -> assert false +let () = define0 "shelve_unifiable" begin + Proofview.shelve_unifiable >>= fun () -> return v_unit +end -let prm_new_goal : ml_tactic = function -| [ev] -> +let () = define1 "new_goal" begin fun ev -> let ev = Evar.unsafe_of_int (Value.to_int ev) in Proofview.tclEVARMAP >>= fun sigma -> if Evd.mem sigma ev then Proofview.Unsafe.tclNEWGOALS [ev] <*> Proofview.tclUNIT v_unit else throw err_notfound -| _ -> assert false +end (** unit -> constr *) -let prm_goal : ml_tactic = function -| [_] -> +let () = define0 "goal" begin Proofview.Goal.enter_one begin fun gl -> let concl = Tacmach.New.pf_nf_concl gl in return (Value.of_constr concl) end -| _ -> assert false +end (** ident -> constr *) -let prm_hyp : ml_tactic = function -| [id] -> +let () = define1 "hyp" begin fun id -> let id = Value.to_ident id in pf_apply begin fun env _ -> let mem = try ignore (Environ.lookup_named id env); true with Not_found -> false in @@ -544,10 +544,9 @@ let prm_hyp : ml_tactic = function else Tacticals.New.tclZEROMSG (str "Hypothesis " ++ quote (Id.print id) ++ str " not found") (** FIXME: Do something more sensible *) end -| _ -> assert false +end -let prm_hyps : ml_tactic = function -| [_] -> +let () = define0 "hyps" begin pf_apply begin fun env _ -> let open Context.Named.Declaration in let hyps = List.rev (Environ.named_context env) in @@ -562,66 +561,59 @@ let prm_hyps : ml_tactic = function in return (Value.of_list map hyps) end -| _ -> assert false +end (** (unit -> constr) -> unit *) -let prm_refine : ml_tactic = function -| [c] -> +let () = define1 "refine" begin fun c -> let c = thaw c >>= fun c -> Proofview.tclUNIT ((), Value.to_constr c) in Proofview.Goal.nf_enter begin fun gl -> Refine.generic_refine ~typecheck:true c gl end >>= fun () -> return v_unit -| _ -> assert false +end -let prm_with_holes : ml_tactic = function -| [x; f] -> +let () = define2 "with_holes" begin fun x f -> Proofview.tclEVARMAP >>= fun sigma0 -> thaw x >>= fun ans -> Proofview.tclEVARMAP >>= fun sigma -> Proofview.Unsafe.tclEVARS sigma0 >>= fun () -> Tacticals.New.tclWITHHOLES false (interp_app f [ans]) sigma -| _ -> assert false +end -let prm_progress : ml_tactic = function -| [f] -> Proofview.tclPROGRESS (thaw f) -| _ -> assert false +let () = define1 "progress" begin fun f -> + Proofview.tclPROGRESS (thaw f) +end -let prm_abstract : ml_tactic = function -| [id; f] -> +let () = define2 "abstract" begin fun id f -> let id = Value.to_option Value.to_ident id in Tactics.tclABSTRACT id (Proofview.tclIGNORE (thaw f)) >>= fun () -> return v_unit -| _ -> assert false +end -let prm_time : ml_tactic = function -| [s; f] -> +let () = define2 "time" begin fun s f -> let s = Value.to_option Value.to_string s in Proofview.tclTIME s (thaw f) -| _ -> assert false +end -let prm_check_interrupt : ml_tactic = function -| [_] -> Proofview.tclCHECKINTERRUPT >>= fun () -> return v_unit -| _ -> assert false +let () = define0 "check_interrupt" begin + Proofview.tclCHECKINTERRUPT >>= fun () -> return v_unit +end (** Fresh *) -let prm_free_union : ml_tactic = function -| [set1; set2] -> +let () = define2 "fresh_free_union" begin fun set1 set2 -> let set1 = Value.to_ext val_free set1 in let set2 = Value.to_ext val_free set2 in let ans = Id.Set.union set1 set2 in return (Value.of_ext val_free ans) -| _ -> assert false +end -let prm_free_of_ids : ml_tactic = function -| [ids] -> +let () = define1 "fresh_free_of_ids" begin fun ids -> let ids = Value.to_list Value.to_ident ids in let free = List.fold_right Id.Set.add ids Id.Set.empty in return (Value.of_ext val_free free) -| _ -> assert false +end -let prm_free_of_constr : ml_tactic = function -| [c] -> +let () = define1 "fresh_free_of_constr" begin fun c -> let c = Value.to_constr c in Proofview.tclEVARMAP >>= fun sigma -> let rec fold accu c = match EConstr.kind sigma c with @@ -630,84 +622,14 @@ let prm_free_of_constr : ml_tactic = function in let ans = fold Id.Set.empty c in return (Value.of_ext val_free ans) -| _ -> assert false +end -let prm_fresh : ml_tactic = function -| [avoid; id] -> +let () = define2 "fresh_fresh" begin fun avoid id -> let avoid = Value.to_ext val_free avoid in let id = Value.to_ident id in let nid = Namegen.next_ident_away_from id (fun id -> Id.Set.mem id avoid) in return (Value.of_ident nid) -| _ -> assert false - -(** Registering *) - -let () = Tac2env.define_primitive (pname "print") prm_print -let () = Tac2env.define_primitive (pname "message_of_string") prm_message_of_string -let () = Tac2env.define_primitive (pname "message_of_int") prm_message_of_int -let () = Tac2env.define_primitive (pname "message_of_constr") prm_message_of_constr -let () = Tac2env.define_primitive (pname "message_of_ident") prm_message_of_ident -let () = Tac2env.define_primitive (pname "message_concat") prm_message_concat - -let () = Tac2env.define_primitive (pname "array_make") prm_array_make -let () = Tac2env.define_primitive (pname "array_length") prm_array_length -let () = Tac2env.define_primitive (pname "array_get") prm_array_get -let () = Tac2env.define_primitive (pname "array_set") prm_array_set - -let () = Tac2env.define_primitive (pname "string_make") prm_string_make -let () = Tac2env.define_primitive (pname "string_length") prm_string_length -let () = Tac2env.define_primitive (pname "string_get") prm_string_get -let () = Tac2env.define_primitive (pname "string_set") prm_string_set - -let () = Tac2env.define_primitive (pname "constr_type") prm_constr_type -let () = Tac2env.define_primitive (pname "constr_equal") prm_constr_equal -let () = Tac2env.define_primitive (pname "constr_kind") prm_constr_kind -let () = Tac2env.define_primitive (pname "constr_substnl") prm_constr_substnl - -let () = Tac2env.define_primitive (pname "pattern_matches") prm_pattern_matches -let () = Tac2env.define_primitive (pname "pattern_matches_subterm") prm_pattern_matches_subterm -let () = Tac2env.define_primitive (pname "pattern_instantiate") prm_pattern_instantiate - -let () = Tac2env.define_primitive (pname "int_equal") prm_int_equal -let () = Tac2env.define_primitive (pname "int_compare") prm_int_compare -let () = Tac2env.define_primitive (pname "int_neg") prm_int_neg -let () = Tac2env.define_primitive (pname "int_add") prm_int_add -let () = Tac2env.define_primitive (pname "int_sub") prm_int_sub -let () = Tac2env.define_primitive (pname "int_mul") prm_int_mul - -let () = Tac2env.define_primitive (pname "ident_equal") prm_ident_equal -let () = Tac2env.define_primitive (pname "ident_to_string") prm_ident_to_string -let () = Tac2env.define_primitive (pname "ident_of_string") prm_ident_of_string - -let () = Tac2env.define_primitive (pname "throw") prm_throw - -let () = Tac2env.define_primitive (pname "zero") prm_zero -let () = Tac2env.define_primitive (pname "plus") prm_plus -let () = Tac2env.define_primitive (pname "once") prm_once -let () = Tac2env.define_primitive (pname "dispatch") prm_dispatch -let () = Tac2env.define_primitive (pname "extend") prm_extend -let () = Tac2env.define_primitive (pname "enter") prm_enter -let () = Tac2env.define_primitive (pname "case") prm_case -let () = Tac2env.define_primitive (pname "apply_kont") prm_apply_kont - -let () = Tac2env.define_primitive (pname "focus") prm_focus -let () = Tac2env.define_primitive (pname "shelve") prm_shelve -let () = Tac2env.define_primitive (pname "shelve_unifiable") prm_shelve_unifiable -let () = Tac2env.define_primitive (pname "new_goal") prm_new_goal -let () = Tac2env.define_primitive (pname "goal") prm_goal -let () = Tac2env.define_primitive (pname "hyp") prm_hyp -let () = Tac2env.define_primitive (pname "hyps") prm_hyps -let () = Tac2env.define_primitive (pname "refine") prm_refine -let () = Tac2env.define_primitive (pname "with_holes") prm_with_holes -let () = Tac2env.define_primitive (pname "progress") prm_progress -let () = Tac2env.define_primitive (pname "abstract") prm_abstract -let () = Tac2env.define_primitive (pname "time") prm_time -let () = Tac2env.define_primitive (pname "check_interrupt") prm_check_interrupt - -let () = Tac2env.define_primitive (pname "fresh_fresh") prm_fresh -let () = Tac2env.define_primitive (pname "fresh_free_union") prm_free_union -let () = Tac2env.define_primitive (pname "fresh_free_of_ids") prm_free_of_ids -let () = Tac2env.define_primitive (pname "fresh_free_of_constr") prm_free_of_constr +end (** ML types *) -- cgit v1.2.3 From f6154c8a086faee725b4f41fb4b2586d7cb6c51d Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 29 Aug 2017 15:19:30 +0200 Subject: Removing dead code for handling of array litterals. --- src/tac2env.ml | 2 +- src/tac2expr.mli | 1 - src/tac2intern.ml | 6 ++---- src/tac2interp.ml | 3 --- src/tac2print.ml | 2 -- 5 files changed, 3 insertions(+), 11 deletions(-) diff --git a/src/tac2env.ml b/src/tac2env.ml index 59344e336b..dd8a07ffc6 100644 --- a/src/tac2env.ml +++ b/src/tac2env.ml @@ -65,7 +65,7 @@ let rec eval_pure = function | GTacCst (_, n, []) -> ValInt n | GTacCst (_, n, el) -> ValBlk (n, Array.map_of_list eval_pure el) | GTacOpn (kn, el) -> ValOpn (kn, Array.map_of_list eval_pure el) -| GTacAtm (AtmStr _) | GTacArr _ | GTacLet _ | GTacVar _ | GTacSet _ +| GTacAtm (AtmStr _) | GTacLet _ | GTacVar _ | GTacSet _ | GTacApp _ | GTacCse _ | GTacPrj _ | GTacPrm _ | GTacExt _ | GTacWth _ -> anomaly (Pp.str "Term is not a syntactical value") diff --git a/src/tac2expr.mli b/src/tac2expr.mli index 0c9112d902..42203a32e5 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -126,7 +126,6 @@ type glb_tacexpr = | GTacFun of Name.t list * glb_tacexpr | GTacApp of glb_tacexpr * glb_tacexpr list | GTacLet of rec_flag * (Name.t * glb_tacexpr) list * glb_tacexpr -| GTacArr of glb_tacexpr list | GTacCst of case_info * int * glb_tacexpr list | GTacCse of glb_tacexpr * case_info * glb_tacexpr array * (Name.t array * glb_tacexpr) array | GTacPrj of type_constant * glb_tacexpr * int diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 9d22b18af4..02dfa1c08b 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -416,13 +416,13 @@ let rec is_value = function | GTacCst (_, _, []) -> true | GTacOpn (_, el) -> List.for_all is_value el | GTacCst (Other kn, _, el) -> is_pure_constructor kn && List.for_all is_value el -| GTacArr _ | GTacCse _ | GTacPrj _ | GTacSet _ | GTacExt _ | GTacPrm _ +| GTacCse _ | GTacPrj _ | GTacSet _ | GTacExt _ | GTacPrm _ | GTacWth _ -> false let is_rec_rhs = function | GTacFun _ -> true | GTacAtm _ | GTacVar _ | GTacRef _ | GTacApp _ | GTacLet _ | GTacPrj _ -| GTacSet _ | GTacArr _ | GTacExt _ | GTacPrm _ | GTacCst _ +| GTacSet _ | GTacExt _ | GTacPrm _ | GTacCst _ | GTacCse _ | GTacOpn _ | GTacWth _ -> false let rec fv_type f t accu = match t with @@ -1294,8 +1294,6 @@ let rec subst_expr subst e = match e with | GTacLet (r, bs, e) -> let bs = List.map (fun (na, e) -> (na, subst_expr subst e)) bs in GTacLet (r, bs, subst_expr subst e) -| GTacArr el -> - GTacArr (List.map (fun e -> subst_expr subst e) el) | GTacCst (t, n, el) as e0 -> let t' = subst_or_tuple subst_kn subst t in let el' = List.smartmap (fun e -> subst_expr subst e) el in diff --git a/src/tac2interp.ml b/src/tac2interp.ml index 3e1a048d29..3490b1a2a8 100644 --- a/src/tac2interp.ml +++ b/src/tac2interp.ml @@ -79,9 +79,6 @@ let rec interp ist = function let iter (_, e) = e.clos_env <- ist in let () = List.iter iter fixs in interp ist e -| GTacArr el -> - Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> - return (ValBlk (0, Array.of_list el)) | GTacCst (_, n, []) -> return (ValInt n) | GTacCst (_, n, el) -> Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> diff --git a/src/tac2print.ml b/src/tac2print.ml index e3095c7a89..29f78f251e 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -177,8 +177,6 @@ let pr_glbexpr_gen lvl c = | E2 | E3 | E4 | E5 -> fun x -> x in paren (prlist_with_sep (fun () -> str "," ++ spc ()) (pr_glbexpr E1) cl) - | GTacArr cl -> - mt () (** FIXME when implemented *) | GTacCst (Other tpe, n, cl) -> begin match Tac2env.interp_type tpe with | _, GTydAlg { galg_constructors = def } -> -- cgit v1.2.3 From 94397f6e022176a29add6369f0a310b1d7decf62 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 29 Aug 2017 15:23:37 +0200 Subject: Pass Ltac2 variables in a dedicated environment for generic arguments. This is a way to hack around the fact that various interpretation functions rely wrongly on the values of the environment to do nasty tricks. Typically, the interpretation of terms is broken, as it will fail when there is a bound variable with the same name as a hypothesis, instead of producing the hypothesis itself. --- src/tac2core.ml | 10 ++-------- src/tac2interp.ml | 19 +++++++++++++++++++ src/tac2interp.mli | 5 +++++ tests/quot.v | 10 ++++++++++ 4 files changed, 36 insertions(+), 8 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 609dd40587..d14849a2a6 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -656,8 +656,7 @@ let open_constr_no_classes_flags () = (** Embed all Ltac2 data into Values *) let to_lvar ist = let open Glob_ops in - let map e = Val.Dyn (val_valexpr, e) in - let lfun = Id.Map.map map ist in + let lfun = Tac2interp.set_env ist Id.Map.empty in { empty_lvar with Glob_term.ltac_genargs = lfun } let interp_constr flags ist (c, _) = @@ -728,12 +727,7 @@ let () = let () = let interp ist env sigma concl tac = - let fold id (Val.Dyn (tag, v)) (accu : environment) : environment = - match Val.eq tag val_valexpr with - | None -> accu - | Some Refl -> Id.Map.add id v accu - in - let ist = Id.Map.fold fold ist Id.Map.empty in + let ist = Tac2interp.get_env ist in let tac = Proofview.tclIGNORE (interp ist tac) in let c, sigma = Pfedit.refine_by_tactic env sigma concl tac in (EConstr.of_constr c, sigma) diff --git a/src/tac2interp.ml b/src/tac2interp.ml index 3490b1a2a8..691c795502 100644 --- a/src/tac2interp.ml +++ b/src/tac2interp.ml @@ -155,3 +155,22 @@ and interp_set ist e p r = match e with return (ValInt 0) | ValInt _ | ValExt _ | ValStr _ | ValCls _ | ValOpn _ -> anomaly (str "Unexpected value shape") + +(** Cross-boundary hacks. *) + +open Geninterp + +let val_env : environment Val.typ = Val.create "ltac2:env" +let env_ref = Id.of_string_soft "@@ltac2_env@@" + +let extract_env (Val.Dyn (tag, v)) : environment = +match Val.eq tag val_env with +| None -> assert false +| Some Refl -> v + +let get_env ist = + try extract_env (Id.Map.find env_ref ist) + with Not_found -> empty_environment + +let set_env env ist = + Id.Map.add env_ref (Val.Dyn (val_env, env)) ist diff --git a/src/tac2interp.mli b/src/tac2interp.mli index 42e9e3adeb..f99008c506 100644 --- a/src/tac2interp.mli +++ b/src/tac2interp.mli @@ -17,6 +17,11 @@ val interp : environment -> glb_tacexpr -> valexpr Proofview.tactic val interp_app : valexpr -> valexpr list -> valexpr Proofview.tactic +(** {5 Cross-boundary encodings} *) + +val get_env : Glob_term.unbound_ltac_var_map -> environment +val set_env : environment -> Glob_term.unbound_ltac_var_map -> Glob_term.unbound_ltac_var_map + (** {5 Exceptions} *) exception LtacError of KerName.t * valexpr array diff --git a/tests/quot.v b/tests/quot.v index 4fa9c4fa4e..624c4ad0c1 100644 --- a/tests/quot.v +++ b/tests/quot.v @@ -14,3 +14,13 @@ let x := constr:(I) in let y := constr:((fun z => z) $x) in Control.refine (fun _ => y). Qed. + +Goal True. +Proof. +(** Here, Ltac2 should not put its variables in the same environment as + Ltac1 otherwise the second binding fails as x is bound but not an + ident. *) +let x := constr:(I) in +let y := constr:((fun x => x) $x) in +Control.refine (fun _ => y). +Qed. -- cgit v1.2.3 From 31e686c2904c3015eaec18ce502d4e8afe565850 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 29 Aug 2017 18:17:19 +0200 Subject: Rolling our own dynamic types for Ltac2. This prevents careless confusions with generic arguments from Coq. --- _CoqProject | 2 ++ src/ltac2_plugin.mlpack | 1 + src/tac2core.ml | 6 +++--- src/tac2dyn.ml | 9 +++++++++ src/tac2dyn.mli | 11 +++++++++++ src/tac2expr.mli | 2 +- src/tac2ffi.ml | 14 +++++--------- src/tac2ffi.mli | 30 +++++++++++++++--------------- src/tac2interp.ml | 2 +- src/tac2interp.mli | 2 +- 10 files changed, 49 insertions(+), 30 deletions(-) create mode 100644 src/tac2dyn.ml create mode 100644 src/tac2dyn.mli diff --git a/_CoqProject b/_CoqProject index ffe1dda032..fc9df4ee3f 100644 --- a/_CoqProject +++ b/_CoqProject @@ -2,6 +2,8 @@ -I src/ -bypass-API +src/tac2dyn.ml +src/tac2dyn.mli src/tac2expr.mli src/tac2env.ml src/tac2env.mli diff --git a/src/ltac2_plugin.mlpack b/src/ltac2_plugin.mlpack index f9fa2fafd8..92f391a085 100644 --- a/src/ltac2_plugin.mlpack +++ b/src/ltac2_plugin.mlpack @@ -1,3 +1,4 @@ +Tac2dyn Tac2env Tac2print Tac2intern diff --git a/src/tac2core.ml b/src/tac2core.ml index d14849a2a6..e865c1378f 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -10,8 +10,8 @@ open CSig open Pp open Names open Genarg -open Geninterp open Tac2env +open Tac2dyn open Tac2expr open Tac2interp open Tac2entries.Pltac @@ -70,8 +70,8 @@ let of_rec_declaration (nas, ts, cs) = Value.of_array Value.of_constr ts, Value.of_array Value.of_constr cs) -let val_valexpr : valexpr Val.typ = Val.create "ltac2:valexpr" -let val_free : Id.Set.t Val.typ = Val.create "ltac2:free" +let val_valexpr : valexpr Val.tag = Val.create "ltac2:valexpr" +let val_free : Id.Set.t Val.tag = Val.create "ltac2:free" (** Stdlib exceptions *) diff --git a/src/tac2dyn.ml b/src/tac2dyn.ml new file mode 100644 index 0000000000..3f4fbca712 --- /dev/null +++ b/src/tac2dyn.ml @@ -0,0 +1,9 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* valexpr +| ValExt : 'a Tac2dyn.Val.tag * 'a -> valexpr (** Arbitrary data *) and closure = { diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index a9a0f5a479..61b6d56b6c 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -9,18 +9,14 @@ open Util open Globnames open Genarg -open Geninterp +open Tac2dyn open Tac2expr open Tac2interp (** Dynamic tags *) -let val_tag t = match val_tag t with -| Val.Base t -> t -| _ -> assert false - -let val_constr = val_tag (topwit Stdarg.wit_constr) -let val_ident = val_tag (topwit Stdarg.wit_ident) +let val_constr = Val.create "ltac2:constr" +let val_ident = Val.create "ltac2:ident" let val_pattern = Val.create "ltac2:pattern" let val_pp = Val.create "ltac2:pp" let val_sort = Val.create "ltac2:sort" @@ -30,10 +26,10 @@ let val_constant = Val.create "ltac2:constant" let val_constructor = Val.create "ltac2:constructor" let val_projection = Val.create "ltac2:projection" let val_univ = Val.create "ltac2:universe" -let val_kont : (Exninfo.iexn -> valexpr Proofview.tactic) Val.typ = +let val_kont : (Exninfo.iexn -> valexpr Proofview.tactic) Val.tag = Val.create "ltac2:kont" -let extract_val (type a) (type b) (tag : a Val.typ) (tag' : b Val.typ) (v : b) : a = +let extract_val (type a) (type b) (tag : a Val.tag) (tag' : b Val.tag) (v : b) : a = match Val.eq tag tag' with | None -> assert false | Some Refl -> v diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index 71d90ba940..1ce163256d 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -6,9 +6,9 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Geninterp open Names open EConstr +open Tac2dyn open Tac2expr (** {5 Ltac2 FFI} *) @@ -65,20 +65,20 @@ val to_constant : valexpr -> Constant.t val of_reference : Globnames.global_reference -> valexpr val to_reference : valexpr -> Globnames.global_reference -val of_ext : 'a Val.typ -> 'a -> valexpr -val to_ext : 'a Val.typ -> valexpr -> 'a +val of_ext : 'a Val.tag -> 'a -> valexpr +val to_ext : 'a Val.tag -> valexpr -> 'a (** {5 Dynamic tags} *) -val val_constr : EConstr.t Val.typ -val val_ident : Id.t Val.typ -val val_pattern : Pattern.constr_pattern Val.typ -val val_pp : Pp.t Val.typ -val val_sort : ESorts.t Val.typ -val val_cast : Constr.cast_kind Val.typ -val val_inductive : inductive Val.typ -val val_constant : Constant.t Val.typ -val val_constructor : constructor Val.typ -val val_projection : Projection.t Val.typ -val val_univ : Univ.universe_level Val.typ -val val_kont : (Exninfo.iexn -> valexpr Proofview.tactic) Val.typ +val val_constr : EConstr.t Val.tag +val val_ident : Id.t Val.tag +val val_pattern : Pattern.constr_pattern Val.tag +val val_pp : Pp.t Val.tag +val val_sort : ESorts.t Val.tag +val val_cast : Constr.cast_kind Val.tag +val val_inductive : inductive Val.tag +val val_constant : Constant.t Val.tag +val val_constructor : constructor Val.tag +val val_projection : Projection.t Val.tag +val val_univ : Univ.universe_level Val.tag +val val_kont : (Exninfo.iexn -> valexpr Proofview.tactic) Val.tag diff --git a/src/tac2interp.ml b/src/tac2interp.ml index 691c795502..3be95ac828 100644 --- a/src/tac2interp.ml +++ b/src/tac2interp.ml @@ -23,7 +23,7 @@ let () = register_handler begin function | _ -> raise Unhandled end -let val_exn = Geninterp.Val.create "ltac2:exn" +let val_exn = Tac2dyn.Val.create "ltac2:exn" type environment = valexpr Id.Map.t diff --git a/src/tac2interp.mli b/src/tac2interp.mli index f99008c506..18522c3c91 100644 --- a/src/tac2interp.mli +++ b/src/tac2interp.mli @@ -27,6 +27,6 @@ val set_env : environment -> Glob_term.unbound_ltac_var_map -> Glob_term.unbound exception LtacError of KerName.t * valexpr array (** Ltac2-defined exceptions seen from OCaml side *) -val val_exn : Exninfo.iexn Geninterp.Val.typ +val val_exn : Exninfo.iexn Tac2dyn.Val.tag (** Toplevel representation of OCaml exceptions. Invariant: no [LtacError] should be put into a value with tag [val_exn]. *) -- cgit v1.2.3 From 63d36d429edd2e85cbebe69f66e8949b25b46c70 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 29 Aug 2017 18:32:38 +0200 Subject: Rolling our own generic arguments. --- src/g_ltac2.ml4 | 6 ++-- src/tac2core.ml | 101 ++++++++++++++++++++++++++++-------------------------- src/tac2dyn.ml | 18 ++++++++++ src/tac2dyn.mli | 23 +++++++++++++ src/tac2env.ml | 35 +++++++++++++------ src/tac2env.mli | 23 +++++++++---- src/tac2expr.mli | 4 +-- src/tac2ffi.ml | 1 + src/tac2ffi.mli | 1 + src/tac2intern.ml | 20 +++++------ src/tac2interp.ml | 3 +- src/tac2print.ml | 10 ++---- src/tac2quote.ml | 8 ++--- 13 files changed, 160 insertions(+), 93 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 672db12f1d..4e62fab715 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -86,11 +86,11 @@ let tac2mode = Gram.entry_create "vernac:ltac2_command" (** FUCK YOU API *) let ltac1_expr = (Obj.magic Pltac.tactic_expr : Tacexpr.raw_tactic_expr Gram.entry) -let inj_wit wit loc x = CTacExt (loc, Genarg.in_gen (Genarg.rawwit wit) x) -let inj_open_constr loc c = inj_wit Stdarg.wit_open_constr loc c +let inj_wit wit loc x = CTacExt (loc, wit, x) +let inj_open_constr loc c = inj_wit Tac2env.wit_open_constr loc c let inj_pattern loc c = inj_wit Tac2env.wit_pattern loc c let inj_reference loc c = inj_wit Tac2env.wit_reference loc c -let inj_ltac1 loc e = inj_wit Tacarg.wit_tactic loc e +let inj_ltac1 loc e = inj_wit Tac2env.wit_ltac1 loc e let pattern_of_qualid loc id = if Tac2env.is_constructor (snd id) then CPatRef (loc, RelId id, []) diff --git a/src/tac2core.ml b/src/tac2core.ml index e865c1378f..cbc7c4744e 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -70,9 +70,6 @@ let of_rec_declaration (nas, ts, cs) = Value.of_array Value.of_constr ts, Value.of_array Value.of_constr cs) -let val_valexpr : valexpr Val.tag = Val.create "ltac2:valexpr" -let val_free : Id.Set.t Val.tag = Val.create "ltac2:free" - (** Stdlib exceptions *) let err_notfocussed = @@ -601,16 +598,16 @@ end (** Fresh *) let () = define2 "fresh_free_union" begin fun set1 set2 -> - let set1 = Value.to_ext val_free set1 in - let set2 = Value.to_ext val_free set2 in + let set1 = Value.to_ext Value.val_free set1 in + let set2 = Value.to_ext Value.val_free set2 in let ans = Id.Set.union set1 set2 in - return (Value.of_ext val_free ans) + return (Value.of_ext Value.val_free ans) end let () = define1 "fresh_free_of_ids" begin fun ids -> let ids = Value.to_list Value.to_ident ids in let free = List.fold_right Id.Set.add ids Id.Set.empty in - return (Value.of_ext val_free free) + return (Value.of_ext Value.val_free free) end let () = define1 "fresh_free_of_constr" begin fun c -> @@ -621,11 +618,11 @@ let () = define1 "fresh_free_of_constr" begin fun c -> | _ -> EConstr.fold sigma fold accu c in let ans = fold Id.Set.empty c in - return (Value.of_ext val_free ans) + return (Value.of_ext Value.val_free ans) end let () = define2 "fresh_fresh" begin fun avoid id -> - let avoid = Value.to_ext val_free avoid in + let avoid = Value.to_ext Value.val_free avoid in let id = Value.to_ident id in let nid = Namegen.next_ident_away_from id (fun id -> Id.Set.mem id avoid) in return (Value.of_ident nid) @@ -659,7 +656,11 @@ let to_lvar ist = let lfun = Tac2interp.set_env ist Id.Map.empty in { empty_lvar with Glob_term.ltac_genargs = lfun } -let interp_constr flags ist (c, _) = +let intern_constr ist c = + let (_, (c, _)) = Genintern.intern Stdarg.wit_constr ist c in + c + +let interp_constr flags ist c = let open Pretyping in pf_apply begin fun env sigma -> Proofview.V82.wrap_exceptions begin fun () -> @@ -672,56 +673,90 @@ let interp_constr flags ist (c, _) = end let () = + let intern = intern_constr in let interp ist c = interp_constr (constr_flags ()) ist c in let obj = { ml_type = t_constr; + ml_intern = intern; + ml_subst = Detyping.subst_glob_constr; ml_interp = interp; } in - define_ml_object Stdarg.wit_constr obj + define_ml_object Tac2env.wit_constr obj let () = + let intern = intern_constr in let interp ist c = interp_constr (open_constr_no_classes_flags ()) ist c in let obj = { ml_type = t_constr; + ml_intern = intern; + ml_subst = Detyping.subst_glob_constr; ml_interp = interp; } in - define_ml_object Stdarg.wit_open_constr obj + define_ml_object Tac2env.wit_open_constr obj let () = let interp _ id = return (ValExt (Value.val_ident, id)) in let obj = { ml_type = t_ident; + ml_intern = (fun _ id -> id); ml_interp = interp; + ml_subst = (fun _ id -> id); } in - define_ml_object Stdarg.wit_ident obj + define_ml_object Tac2env.wit_ident obj let () = + let intern ist c = + let _, pat = Constrintern.intern_constr_pattern ist.Genintern.genv ~as_type:false c in + pat + in let interp _ c = return (ValExt (Value.val_pattern, c)) in let obj = { ml_type = t_pattern; + ml_intern = intern; ml_interp = interp; + ml_subst = Patternops.subst_pattern; } in define_ml_object Tac2env.wit_pattern obj let () = + let intern ist qid = match qid with + | Libnames.Ident (_, id) -> Globnames.VarRef id + | Libnames.Qualid (loc, qid) -> + let gr = + try Nametab.locate qid + with Not_found -> + Nametab.error_global_not_found ?loc qid + in + gr + in + let subst s c = Globnames.subst_global_reference s c in let interp _ gr = return (Value.of_reference gr) in let obj = { ml_type = t_reference; + ml_intern = intern; + ml_subst = subst; ml_interp = interp; } in define_ml_object Tac2env.wit_reference obj let () = + let intern ist tac = + let _, tac = Genintern.intern Ltac_plugin.Tacarg.wit_tactic ist tac in + tac + in let interp _ tac = (** FUCK YOU API *) (Obj.magic Ltac_plugin.Tacinterp.eval_tactic tac : unit Proofview.tactic) >>= fun () -> return v_unit in + let subst s tac = Genintern.substitute Ltac_plugin.Tacarg.wit_tactic s tac in let obj = { ml_type = t_unit; + ml_intern = intern; + ml_subst = subst; ml_interp = interp; } in - define_ml_object Ltac_plugin.Tacarg.wit_tactic obj + define_ml_object Tac2env.wit_ltac1 obj (** Ltac2 in terms *) @@ -754,38 +789,6 @@ let () = in Geninterp.register_interp0 wit_ltac2 interp -(** Patterns *) - -let () = - let intern ist c = - let _, pat = Constrintern.intern_constr_pattern ist.Genintern.genv ~as_type:false c in - ist, pat - in - Genintern.register_intern0 wit_pattern intern - -let () = - let subst s c = Patternops.subst_pattern s c in - Genintern.register_subst0 wit_pattern subst - -(** References *) - -let () = - let intern ist qid = match qid with - | Libnames.Ident (_, id) -> ist, Globnames.VarRef id - | Libnames.Qualid (loc, qid) -> - let gr = - try Nametab.locate qid - with Not_found -> - Nametab.error_global_not_found ?loc qid - in - ist, gr - in - Genintern.register_intern0 wit_reference intern - -let () = - let subst s c = Globnames.subst_global_reference s c in - Genintern.register_subst0 wit_reference subst - (** Built-in notation scopes *) let add_scope s f = @@ -806,7 +809,7 @@ let add_generic_scope s entry arg = let parse = function | [] -> let scope = Extend.Aentry entry in - let act x = CTacExt (dummy_loc, in_gen (rawwit arg) x) in + let act x = CTacExt (dummy_loc, arg, x) in Tac2entries.ScopeRule (scope, act) | _ -> scope_fail () in @@ -927,8 +930,8 @@ let () = add_expr_scope "dispatch" q_dispatch Tac2quote.of_dispatch let () = add_expr_scope "strategy" q_strategy_flag Tac2quote.of_strategy_flag let () = add_expr_scope "reference" q_reference Tac2quote.of_reference -let () = add_generic_scope "constr" Pcoq.Constr.constr Stdarg.wit_constr -let () = add_generic_scope "open_constr" Pcoq.Constr.constr Stdarg.wit_open_constr +let () = add_generic_scope "constr" Pcoq.Constr.constr wit_constr +let () = add_generic_scope "open_constr" Pcoq.Constr.constr wit_open_constr let () = add_generic_scope "pattern" Pcoq.Constr.constr wit_pattern (** seq scope, a bit hairy *) diff --git a/src/tac2dyn.ml b/src/tac2dyn.ml index 3f4fbca712..896676f08b 100644 --- a/src/tac2dyn.ml +++ b/src/tac2dyn.ml @@ -6,4 +6,22 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +module Arg = +struct + module DYN = Dyn.Make(struct end) + module Map = DYN.Map + type ('a, 'b) tag = ('a * 'b) DYN.tag + let eq = DYN.eq + let repr = DYN.repr + let create = DYN.create +end + +module type Param = sig type ('raw, 'glb) t end + +module ArgMap (M : Param) = +struct + type _ pack = Pack : ('raw, 'glb) M.t -> ('raw * 'glb) pack + include Arg.Map(struct type 'a t = 'a pack end) +end + module Val = Dyn.Make(struct end) diff --git a/src/tac2dyn.mli b/src/tac2dyn.mli index e4b18ba373..e995296840 100644 --- a/src/tac2dyn.mli +++ b/src/tac2dyn.mli @@ -8,4 +8,27 @@ (** Dynamic arguments for Ltac2. *) +module Arg : +sig + type ('a, 'b) tag + val create : string -> ('a, 'b) tag + val eq : ('a1, 'b1) tag -> ('a2, 'b2) tag -> ('a1 * 'b1, 'a2 * 'b2) CSig.eq option + val repr : ('a, 'b) tag -> string +end +(** Arguments that are part of an AST. *) + +module type Param = sig type ('raw, 'glb) t end + +module ArgMap (M : Param) : +sig + type _ pack = Pack : ('raw, 'glb) M.t -> ('raw * 'glb) pack + type t + val empty : t + val add : ('a, 'b) Arg.tag -> ('a * 'b) pack -> t -> t + val remove : ('a, 'b) Arg.tag -> t -> t + val find : ('a, 'b) Arg.tag -> t -> ('a * 'b) pack + val mem : ('a, 'b) Arg.tag -> t -> bool +end + module Val : Dyn.S +(** Toplevel values *) diff --git a/src/tac2env.ml b/src/tac2env.ml index dd8a07ffc6..39821b1cb6 100644 --- a/src/tac2env.ml +++ b/src/tac2env.ml @@ -10,6 +10,7 @@ open CErrors open Util open Names open Libnames +open Tac2dyn open Tac2expr type global_data = { @@ -250,22 +251,31 @@ let shortest_qualid_of_projection kn = let sp = KNmap.find kn tab.tab_proj_rev in KnTab.shortest_qualid Id.Set.empty sp tab.tab_proj -type 'a ml_object = { +type ('a, 'b) ml_object = { ml_type : type_constant; - ml_interp : environment -> 'a -> valexpr Proofview.tactic; + ml_intern : Genintern.glob_sign -> 'a -> 'b; + ml_subst : Mod_subst.substitution -> 'b -> 'b; + ml_interp : environment -> 'b -> valexpr Proofview.tactic; } module MLTypeObj = struct - type ('a, 'b, 'c) obj = 'b ml_object - let name = "ltac2_ml_type" - let default _ = None + type ('a, 'b) t = ('a, 'b) ml_object end -module MLType = Genarg.Register(MLTypeObj) +module MLType = Tac2dyn.ArgMap(MLTypeObj) -let define_ml_object t tpe = MLType.register0 t tpe -let interp_ml_object t = MLType.obj t +let ml_object_table = ref MLType.empty + +let define_ml_object t tpe = + ml_object_table := MLType.add t (MLType.Pack tpe) !ml_object_table + +let interp_ml_object t = + try + let MLType.Pack ans = MLType.find t !ml_object_table in + ans + with Not_found -> + CErrors.anomaly Pp.(str "Unknown object type " ++ str (Tac2dyn.Arg.repr t)) (** Absolute paths *) @@ -278,8 +288,13 @@ let std_prefix = (** Generic arguments *) let wit_ltac2 = Genarg.make0 "ltac2:value" -let wit_pattern = Genarg.make0 "ltac2:pattern" -let wit_reference = Genarg.make0 "ltac2:reference" + +let wit_pattern = Arg.create "pattern" +let wit_reference = Arg.create "reference" +let wit_ident = Arg.create "ident" +let wit_constr = Arg.create "constr" +let wit_open_constr = Arg.create "open_constr" +let wit_ltac1 = Arg.create "ltac1" let is_constructor qid = let (_, id) = repr_qualid qid in diff --git a/src/tac2env.mli b/src/tac2env.mli index 8a5fb531d8..0ab6543178 100644 --- a/src/tac2env.mli +++ b/src/tac2env.mli @@ -10,6 +10,7 @@ open Genarg open Names open Libnames open Nametab +open Tac2dyn open Tac2expr (** Ltac2 global environment *) @@ -104,13 +105,15 @@ val interp_primitive : ml_tactic_name -> ml_tactic (** {5 ML primitive types} *) -type 'a ml_object = { +type ('a, 'b) ml_object = { ml_type : type_constant; - ml_interp : environment -> 'a -> valexpr Proofview.tactic; + ml_intern : Genintern.glob_sign -> 'a -> 'b; + ml_subst : Mod_subst.substitution -> 'b -> 'b; + ml_interp : environment -> 'b -> valexpr Proofview.tactic; } -val define_ml_object : ('a, 'b, 'c) genarg_type -> 'b ml_object -> unit -val interp_ml_object : ('a, 'b, 'c) genarg_type -> 'b ml_object +val define_ml_object : ('a, 'b) Tac2dyn.Arg.tag -> ('a, 'b) ml_object -> unit +val interp_ml_object : ('a, 'b) Tac2dyn.Arg.tag -> ('a, 'b) ml_object (** {5 Absolute paths} *) @@ -124,13 +127,21 @@ val std_prefix : ModPath.t val wit_ltac2 : (raw_tacexpr, glb_tacexpr, Util.Empty.t) genarg_type -val wit_pattern : (Constrexpr.constr_expr, Pattern.constr_pattern, Util.Empty.t) genarg_type +val wit_pattern : (Constrexpr.constr_expr, Pattern.constr_pattern) Arg.tag -val wit_reference : (reference, Globnames.global_reference, Util.Empty.t) genarg_type +val wit_ident : (Id.t, Id.t) Arg.tag + +val wit_reference : (reference, Globnames.global_reference) Arg.tag (** Beware, at the raw level, [Qualid [id]] has not the same meaning as [Ident id]. The first is an unqualified global reference, the second is the dynamic reference to id. *) +val wit_constr : (Constrexpr.constr_expr, Glob_term.glob_constr) Arg.tag + +val wit_open_constr : (Constrexpr.constr_expr, Glob_term.glob_constr) Arg.tag + +val wit_ltac1 : (Tacexpr.raw_tactic_expr, Tacexpr.glob_tactic_expr) Arg.tag + (** {5 Helper functions} *) val is_constructor : qualid -> bool diff --git a/src/tac2expr.mli b/src/tac2expr.mli index 0b02ba2656..ccff8e7756 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -104,7 +104,7 @@ type raw_tacexpr = | CTacRec of Loc.t * raw_recexpr | CTacPrj of Loc.t * raw_tacexpr * ltac_projection or_relid | CTacSet of Loc.t * raw_tacexpr * ltac_projection or_relid * raw_tacexpr -| CTacExt of Loc.t * raw_generic_argument +| CTacExt : Loc.t * ('a, _) Tac2dyn.Arg.tag * 'a -> raw_tacexpr and raw_taccase = raw_patexpr * raw_tacexpr @@ -132,7 +132,7 @@ type glb_tacexpr = | GTacSet of type_constant * glb_tacexpr * int * glb_tacexpr | GTacOpn of ltac_constructor * glb_tacexpr list | GTacWth of glb_tacexpr open_match -| GTacExt of glob_generic_argument +| GTacExt : (_, 'a) Tac2dyn.Arg.tag * 'a -> glb_tacexpr | GTacPrm of ml_tactic_name * glb_tacexpr list (** {5 Parsing & Printing} *) diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index 61b6d56b6c..6fc3b2e0f2 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -28,6 +28,7 @@ let val_projection = Val.create "ltac2:projection" let val_univ = Val.create "ltac2:universe" let val_kont : (Exninfo.iexn -> valexpr Proofview.tactic) Val.tag = Val.create "ltac2:kont" +let val_free : Names.Id.Set.t Val.tag = Val.create "ltac2:free" let extract_val (type a) (type b) (tag : a Val.tag) (tag' : b Val.tag) (v : b) : a = match Val.eq tag tag' with diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index 1ce163256d..33b1010213 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -82,3 +82,4 @@ val val_constructor : constructor Val.tag val val_projection : Projection.t Val.tag val val_univ : Univ.universe_level Val.tag val val_kont : (Exninfo.iexn -> valexpr Proofview.tactic) Val.tag +val val_free : Id.Set.t Val.tag diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 02dfa1c08b..051b3af5c7 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -200,7 +200,7 @@ let loc_of_tacexpr = function | CTacRec (loc, _) -> loc | CTacPrj (loc, _, _) -> loc | CTacSet (loc, _, _, _) -> loc -| CTacExt (loc, _) -> loc +| CTacExt (loc, _, _) -> loc let loc_of_patexpr = function | CPatVar (loc, _) -> Option.default dummy_loc loc @@ -769,17 +769,16 @@ let rec intern_rec env = function let ret = subst_type substf pinfo.pdata_ptyp in let r = intern_rec_with_constraint env r ret in (GTacSet (pinfo.pdata_type, e, pinfo.pdata_indx, r), GTypRef (Tuple 0, [])) -| CTacExt (loc, ext) -> +| CTacExt (loc, tag, arg) -> let open Genintern in - let GenArg (Rawwit tag, _) = ext in let tpe = interp_ml_object tag in (** External objects do not have access to the named context because this is not stable by dynamic semantics. *) let genv = Global.env_of_context Environ.empty_named_context_val in let ist = empty_glob_sign genv in let ist = { ist with extra = Store.set ist.extra ltac2_env env } in - let (_, ext) = Flags.with_option Ltac_plugin.Tacintern.strict_check (fun () -> generic_intern ist ext) () in - (GTacExt ext, GTypRef (Other tpe.ml_type, [])) + let arg = Flags.with_option Ltac_plugin.Tacintern.strict_check (fun () -> tpe.ml_intern ist arg) () in + (GTacExt (tag, arg), GTypRef (Other tpe.ml_type, [])) and intern_rec_with_constraint env e exp = let loc = loc_of_tacexpr e in @@ -1248,8 +1247,8 @@ let rec globalize ids e = match e with let p = get_projection0 p in let e' = globalize ids e' in CTacSet (loc, e, AbsKn p, e') -| CTacExt (loc, arg) -> - let arg = pr_argument_type (genarg_tag arg) in +| CTacExt (loc, tag, arg) -> + let arg = str (Tac2dyn.Arg.repr tag) in CErrors.user_err ~loc (str "Cannot globalize generic arguments of type" ++ spc () ++ arg) and globalize_case ids (p, e) = @@ -1324,9 +1323,10 @@ let rec subst_expr subst e = match e with let e' = subst_expr subst e in let r' = subst_expr subst r in if kn' == kn && e' == e && r' == r then e0 else GTacSet (kn', e', p, r') -| GTacExt ext -> - let ext' = Genintern.generic_substitute subst ext in - if ext' == ext then e else GTacExt ext' +| GTacExt (tag, arg) -> + let tpe = interp_ml_object tag in + let arg' = tpe.ml_subst subst arg in + if arg' == arg then e else GTacExt (tag, arg') | GTacOpn (kn, el) as e0 -> let kn' = subst_kn subst kn in let el' = List.smartmap (fun e -> subst_expr subst e) el in diff --git a/src/tac2interp.ml b/src/tac2interp.ml index 3be95ac828..f458f6e81f 100644 --- a/src/tac2interp.ml +++ b/src/tac2interp.ml @@ -99,8 +99,7 @@ let rec interp ist = function | GTacPrm (ml, el) -> Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> Tac2env.interp_primitive ml el -| GTacExt e -> - let GenArg (Glbwit tag, e) = e in +| GTacExt (tag, e) -> let tpe = Tac2env.interp_ml_object tag in tpe.Tac2env.ml_interp ist e diff --git a/src/tac2print.ml b/src/tac2print.ml index 29f78f251e..6943697b45 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -279,13 +279,9 @@ let pr_glbexpr_gen lvl c = in let c = pr_constructor kn in paren (hov 0 (c ++ spc () ++ (pr_sequence (pr_glbexpr E0) cl))) - | GTacExt arg -> - let GenArg (Glbwit tag, arg) = arg in - let name = match tag with - | ExtraArg tag -> ArgT.repr tag - | _ -> assert false - in - hov 0 (str name ++ str ":" ++ paren (Genprint.glb_print tag arg)) + | GTacExt (tag, arg) -> + let name = Tac2dyn.Arg.repr tag in + hov 0 (str name ++ str ":" ++ paren (str "_")) (** FIXME *) | GTacPrm (prm, args) -> let args = match args with | [] -> mt () diff --git a/src/tac2quote.ml b/src/tac2quote.ml index f87e370032..279ab53b67 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -65,7 +65,7 @@ let of_option ?loc f opt = match opt with let inj_wit ?loc wit x = let loc = Option.default dummy_loc loc in - CTacExt (loc, Genarg.in_gen (Genarg.rawwit wit) x) + CTacExt (loc, wit, x) let of_variable (loc, id) = let qid = Libnames.qualid_of_ident id in @@ -77,15 +77,15 @@ let of_anti f = function | QExpr x -> f x | QAnti id -> of_variable id -let of_ident (loc, id) = inj_wit ?loc Stdarg.wit_ident id +let of_ident (loc, id) = inj_wit ?loc Tac2env.wit_ident id let of_constr c = let loc = Constrexpr_ops.constr_loc c in - inj_wit ?loc Stdarg.wit_constr c + inj_wit ?loc Tac2env.wit_constr c let of_open_constr c = let loc = Constrexpr_ops.constr_loc c in - inj_wit ?loc Stdarg.wit_open_constr c + inj_wit ?loc Tac2env.wit_open_constr c let of_bool ?loc b = let c = if b then coq_core "true" else coq_core "false" in -- cgit v1.2.3 From 033ac67e2513f2bd3588cc577906538f5b291da4 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 29 Aug 2017 21:18:36 +0200 Subject: Centralizing tag declarations. --- src/tac2ffi.ml | 27 ++++++++++++++------------- src/tac2ffi.mli | 4 ++++ src/tac2interp.ml | 2 -- src/tac2interp.mli | 4 ---- 4 files changed, 18 insertions(+), 19 deletions(-) diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index 6fc3b2e0f2..4d84da14ce 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -15,20 +15,21 @@ open Tac2interp (** Dynamic tags *) -let val_constr = Val.create "ltac2:constr" -let val_ident = Val.create "ltac2:ident" -let val_pattern = Val.create "ltac2:pattern" -let val_pp = Val.create "ltac2:pp" -let val_sort = Val.create "ltac2:sort" -let val_cast = Val.create "ltac2:cast" -let val_inductive = Val.create "ltac2:inductive" -let val_constant = Val.create "ltac2:constant" -let val_constructor = Val.create "ltac2:constructor" -let val_projection = Val.create "ltac2:projection" -let val_univ = Val.create "ltac2:universe" +let val_exn = Val.create "exn" +let val_constr = Val.create "constr" +let val_ident = Val.create "ident" +let val_pattern = Val.create "pattern" +let val_pp = Val.create "pp" +let val_sort = Val.create "sort" +let val_cast = Val.create "cast" +let val_inductive = Val.create "inductive" +let val_constant = Val.create "constant" +let val_constructor = Val.create "constructor" +let val_projection = Val.create "projection" +let val_univ = Val.create "universe" let val_kont : (Exninfo.iexn -> valexpr Proofview.tactic) Val.tag = - Val.create "ltac2:kont" -let val_free : Names.Id.Set.t Val.tag = Val.create "ltac2:free" + Val.create "kont" +let val_free : Names.Id.Set.t Val.tag = Val.create "free" let extract_val (type a) (type b) (tag : a Val.tag) (tag' : b Val.tag) (v : b) : a = match Val.eq tag tag' with diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index 33b1010213..cf1d7e81a1 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -83,3 +83,7 @@ val val_projection : Projection.t Val.tag val val_univ : Univ.universe_level Val.tag val val_kont : (Exninfo.iexn -> valexpr Proofview.tactic) Val.tag val val_free : Id.Set.t Val.tag + +val val_exn : Exninfo.iexn Tac2dyn.Val.tag +(** Toplevel representation of OCaml exceptions. Invariant: no [LtacError] + should be put into a value with tag [val_exn]. *) diff --git a/src/tac2interp.ml b/src/tac2interp.ml index f458f6e81f..c15331571b 100644 --- a/src/tac2interp.ml +++ b/src/tac2interp.ml @@ -23,8 +23,6 @@ let () = register_handler begin function | _ -> raise Unhandled end -let val_exn = Tac2dyn.Val.create "ltac2:exn" - type environment = valexpr Id.Map.t let empty_environment = Id.Map.empty diff --git a/src/tac2interp.mli b/src/tac2interp.mli index 18522c3c91..1ac26b48db 100644 --- a/src/tac2interp.mli +++ b/src/tac2interp.mli @@ -26,7 +26,3 @@ val set_env : environment -> Glob_term.unbound_ltac_var_map -> Glob_term.unbound exception LtacError of KerName.t * valexpr array (** Ltac2-defined exceptions seen from OCaml side *) - -val val_exn : Exninfo.iexn Tac2dyn.Val.tag -(** Toplevel representation of OCaml exceptions. Invariant: no [LtacError] - should be put into a value with tag [val_exn]. *) -- cgit v1.2.3 From ba68fcd85dd38f0094c8eac157080670354e473e Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 29 Aug 2017 22:31:34 +0200 Subject: Fixing printing of tactic expressions. --- src/tac2core.ml | 17 +++++++++++++++++ src/tac2env.ml | 1 + src/tac2env.mli | 1 + src/tac2print.ml | 4 ++-- 4 files changed, 21 insertions(+), 2 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index cbc7c4744e..7d18bf693e 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -675,32 +675,38 @@ let interp_constr flags ist c = let () = let intern = intern_constr in let interp ist c = interp_constr (constr_flags ()) ist c in + let print env c = str "constr:(" ++ Printer.pr_lglob_constr_env env c ++ str ")" in let obj = { ml_type = t_constr; ml_intern = intern; ml_subst = Detyping.subst_glob_constr; ml_interp = interp; + ml_print = print; } in define_ml_object Tac2env.wit_constr obj let () = let intern = intern_constr in let interp ist c = interp_constr (open_constr_no_classes_flags ()) ist c in + let print env c = str "open_constr:(" ++ Printer.pr_lglob_constr_env env c ++ str ")" in let obj = { ml_type = t_constr; ml_intern = intern; ml_subst = Detyping.subst_glob_constr; ml_interp = interp; + ml_print = print; } in define_ml_object Tac2env.wit_open_constr obj let () = let interp _ id = return (ValExt (Value.val_ident, id)) in + let print _ id = str "ident:(" ++ Id.print id ++ str ")" in let obj = { ml_type = t_ident; ml_intern = (fun _ id -> id); ml_interp = interp; ml_subst = (fun _ id -> id); + ml_print = print; } in define_ml_object Tac2env.wit_ident obj @@ -709,12 +715,14 @@ let () = let _, pat = Constrintern.intern_constr_pattern ist.Genintern.genv ~as_type:false c in pat in + let print env pat = str "pattern:(" ++ Printer.pr_lconstr_pattern_env env Evd.empty pat ++ str ")" in let interp _ c = return (ValExt (Value.val_pattern, c)) in let obj = { ml_type = t_pattern; ml_intern = intern; ml_interp = interp; ml_subst = Patternops.subst_pattern; + ml_print = print; } in define_ml_object Tac2env.wit_pattern obj @@ -731,11 +739,16 @@ let () = in let subst s c = Globnames.subst_global_reference s c in let interp _ gr = return (Value.of_reference gr) in + let print _ = function + | Globnames.VarRef id -> str "reference:(" ++ str "&" ++ Id.print id ++ str ")" + | r -> str "reference:(" ++ Printer.pr_global r ++ str ")" + in let obj = { ml_type = t_reference; ml_intern = intern; ml_subst = subst; ml_interp = interp; + ml_print = print; } in define_ml_object Tac2env.wit_reference obj @@ -750,11 +763,15 @@ let () = return v_unit in let subst s tac = Genintern.substitute Ltac_plugin.Tacarg.wit_tactic s tac in + let print env tac = + str "ltac1:(" ++ Ltac_plugin.Pptactic.pr_glob_tactic (Obj.magic env) tac ++ str ")" + in let obj = { ml_type = t_unit; ml_intern = intern; ml_subst = subst; ml_interp = interp; + ml_print = print; } in define_ml_object Tac2env.wit_ltac1 obj diff --git a/src/tac2env.ml b/src/tac2env.ml index 39821b1cb6..489113c031 100644 --- a/src/tac2env.ml +++ b/src/tac2env.ml @@ -256,6 +256,7 @@ type ('a, 'b) ml_object = { ml_intern : Genintern.glob_sign -> 'a -> 'b; ml_subst : Mod_subst.substitution -> 'b -> 'b; ml_interp : environment -> 'b -> valexpr Proofview.tactic; + ml_print : Environ.env -> 'b -> Pp.t; } module MLTypeObj = diff --git a/src/tac2env.mli b/src/tac2env.mli index 0ab6543178..0ef62d67ed 100644 --- a/src/tac2env.mli +++ b/src/tac2env.mli @@ -110,6 +110,7 @@ type ('a, 'b) ml_object = { ml_intern : Genintern.glob_sign -> 'a -> 'b; ml_subst : Mod_subst.substitution -> 'b -> 'b; ml_interp : environment -> 'b -> valexpr Proofview.tactic; + ml_print : Environ.env -> 'b -> Pp.t; } val define_ml_object : ('a, 'b) Tac2dyn.Arg.tag -> ('a, 'b) ml_object -> unit diff --git a/src/tac2print.ml b/src/tac2print.ml index 6943697b45..939f44aeaf 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -280,8 +280,8 @@ let pr_glbexpr_gen lvl c = let c = pr_constructor kn in paren (hov 0 (c ++ spc () ++ (pr_sequence (pr_glbexpr E0) cl))) | GTacExt (tag, arg) -> - let name = Tac2dyn.Arg.repr tag in - hov 0 (str name ++ str ":" ++ paren (str "_")) (** FIXME *) + let tpe = interp_ml_object tag in + hov 0 (tpe.ml_print (Global.env ()) arg) (** FIXME *) | GTacPrm (prm, args) -> let args = match args with | [] -> mt () -- cgit v1.2.3 From 93e888000664191fa608a8fa0f8057bdda8fe084 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 29 Aug 2017 22:51:37 +0200 Subject: Fix printing of Ltac2 in quotations. --- src/tac2core.ml | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/src/tac2core.ml b/src/tac2core.ml index 7d18bf693e..95fd29ec33 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -806,6 +806,12 @@ let () = in Geninterp.register_interp0 wit_ltac2 interp +let () = + let pr_raw _ = mt () in + let pr_glb e = Tac2print.pr_glbexpr e in + let pr_top _ = mt () in + Genprint.register_print0 wit_ltac2 pr_raw pr_glb pr_top + (** Built-in notation scopes *) let add_scope s f = -- cgit v1.2.3 From 5a157fdc706860473638b295c95dd2a6eaa33a41 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 29 Aug 2017 23:23:26 +0200 Subject: Binding reduction functions acting on terms. --- src/tac2core.mli | 2 ++ src/tac2stdlib.ml | 89 +++++++++++++++++++++++++++++++++++++++++++++++++++++ src/tac2tactics.ml | 57 ++++++++++++++++++++++++++++++++++ src/tac2tactics.mli | 24 +++++++++++++++ theories/Std.v | 13 ++++++++ 5 files changed, 185 insertions(+) diff --git a/src/tac2core.mli b/src/tac2core.mli index 566b7efbb7..9fae65bb3e 100644 --- a/src/tac2core.mli +++ b/src/tac2core.mli @@ -26,3 +26,5 @@ val c_true : ltac_constructor val c_false : ltac_constructor end + +val pf_apply : (Environ.env -> Evd.evar_map -> 'a Proofview.tactic) -> 'a Proofview.tactic diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index d88402cbf2..2a57fdc705 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -352,6 +352,95 @@ let () = define_prim2 "tac_native" begin fun where cl -> Tac2tactics.native where cl end +(** Reduction functions *) + +let define_red1 name tac = + let tac = function + | [x] -> tac x >>= fun c -> Proofview.tclUNIT (Value.of_constr c) + | _ -> assert false + in + Tac2env.define_primitive (pname name) tac + +let define_red2 name tac = + let tac = function + | [x; y] -> tac x y >>= fun c -> Proofview.tclUNIT (Value.of_constr c) + | _ -> assert false + in + Tac2env.define_primitive (pname name) tac + +let define_red3 name tac = + let tac = function + | [x; y; z] -> tac x y z >>= fun c -> Proofview.tclUNIT (Value.of_constr c) + | _ -> assert false + in + Tac2env.define_primitive (pname name) tac + +let () = define_red1 "eval_red" begin fun c -> + let c = Value.to_constr c in + Tac2tactics.eval_red c +end + +let () = define_red1 "eval_hnf" begin fun c -> + let c = Value.to_constr c in + Tac2tactics.eval_hnf c +end + +let () = define_red3 "eval_simpl" begin fun flags where c -> + let flags = to_red_flag flags in + let where = Value.to_option to_pattern_with_occs where in + let c = Value.to_constr c in + Tac2tactics.eval_simpl flags where c +end + +let () = define_red2 "eval_cbv" begin fun flags c -> + let flags = to_red_flag flags in + let c = Value.to_constr c in + Tac2tactics.eval_cbv flags c +end + +let () = define_red2 "eval_cbn" begin fun flags c -> + let flags = to_red_flag flags in + let c = Value.to_constr c in + Tac2tactics.eval_cbn flags c +end + +let () = define_red2 "eval_lazy" begin fun flags c -> + let flags = to_red_flag flags in + let c = Value.to_constr c in + Tac2tactics.eval_lazy flags c +end + +let () = define_red2 "eval_unfold" begin fun refs c -> + let map v = to_pair Value.to_reference (fun occ -> to_occurrences to_int_or_var occ) v in + let refs = Value.to_list map refs in + let c = Value.to_constr c in + Tac2tactics.eval_unfold refs c +end + +let () = define_red2 "eval_fold" begin fun args c -> + let args = Value.to_list Value.to_constr args in + let c = Value.to_constr c in + Tac2tactics.eval_fold args c +end + +let () = define_red2 "eval_pattern" begin fun where c -> + let where = Value.to_list (fun p -> to_pair Value.to_constr (fun occ -> to_occurrences to_int_or_var occ) p) where in + let c = Value.to_constr c in + Tac2tactics.eval_pattern where c +end + +let () = define_red2 "eval_vm" begin fun where c -> + let where = Value.to_option to_pattern_with_occs where in + let c = Value.to_constr c in + Tac2tactics.eval_vm where c +end + +let () = define_red2 "eval_native" begin fun where c -> + let where = Value.to_option to_pattern_with_occs where in + let c = Value.to_constr c in + Tac2tactics.eval_native where c +end + let () = define_prim4 "tac_rewrite" begin fun ev rw cl by -> let ev = Value.to_bool ev in let rw = Value.to_list to_rewriting rw in diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index 25a00fdc2e..a95e60bc97 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -122,6 +122,63 @@ let native where cl = let where = Option.map map_pattern_with_occs where in Tactics.reduce (CbvNative where) cl +let eval_fun red c = + Tac2core.pf_apply begin fun env sigma -> + let (redfun, _) = Redexpr.reduction_of_red_expr env red in + let (sigma, ans) = redfun env sigma c in + Proofview.Unsafe.tclEVARS sigma >>= fun () -> + Proofview.tclUNIT ans + end + +let eval_red c = + eval_fun (Red false) c + +let eval_hnf c = + eval_fun Hnf c + +let eval_simpl flags where c = + let where = Option.map map_pattern_with_occs where in + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + eval_fun (Simpl (flags, where)) c + +let eval_cbv flags c = + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + eval_fun (Cbv flags) c + +let eval_cbn flags c = + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + eval_fun (Cbn flags) c + +let eval_lazy flags c = + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + eval_fun (Lazy flags) c + +let eval_unfold occs c = + let map (gr, occ) = + get_evaluable_reference gr >>= fun gr -> Proofview.tclUNIT (occ, gr) + in + Proofview.Monad.List.map map occs >>= fun occs -> + eval_fun (Unfold occs) c + +let eval_fold cl c = + eval_fun (Fold cl) c + +let eval_pattern where c = + let where = List.map (fun (pat, occ) -> (occ, pat)) where in + eval_fun (Pattern where) c + +let eval_vm where c = + let where = Option.map map_pattern_with_occs where in + eval_fun (CbvVm where) c + +let eval_native where c = + let where = Option.map map_pattern_with_occs where in + eval_fun (CbvNative where) c + let on_destruction_arg tac ev arg = Proofview.Goal.enter begin fun gl -> match arg with diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli index 8939d2a9dd..87489bb248 100644 --- a/src/tac2tactics.mli +++ b/src/tac2tactics.mli @@ -9,6 +9,7 @@ open Names open Locus open Globnames +open EConstr open Genredexpr open Misctypes open Tactypes @@ -54,6 +55,29 @@ val vm : (Pattern.constr_pattern * occurrences_expr) option -> clause -> unit ta val native : (Pattern.constr_pattern * occurrences_expr) option -> clause -> unit tactic +val eval_red : constr -> constr tactic + +val eval_hnf : constr -> constr tactic + +val eval_simpl : global_reference glob_red_flag -> + (Pattern.constr_pattern * occurrences_expr) option -> constr -> constr tactic + +val eval_cbv : global_reference glob_red_flag -> constr -> constr tactic + +val eval_cbn : global_reference glob_red_flag -> constr -> constr tactic + +val eval_lazy : global_reference glob_red_flag -> constr -> constr tactic + +val eval_unfold : (global_reference * occurrences_expr) list -> constr -> constr tactic + +val eval_fold : constr list -> constr -> constr tactic + +val eval_pattern : (EConstr.t * occurrences_expr) list -> constr -> constr tactic + +val eval_vm : (Pattern.constr_pattern * occurrences_expr) option -> constr -> constr tactic + +val eval_native : (Pattern.constr_pattern * occurrences_expr) option -> constr -> constr tactic + val discriminate : evars_flag -> destruction_arg option -> unit tactic val injection : evars_flag -> intro_pattern list option -> destruction_arg option -> unit tactic diff --git a/theories/Std.v b/theories/Std.v index f380c10af8..2fa2c34da6 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -142,6 +142,19 @@ Ltac2 @ external pattern : (constr * occurrences) list -> clause -> unit := "lta Ltac2 @ external vm : (pattern * occurrences) option -> clause -> unit := "ltac2" "tac_vm". Ltac2 @ external native : (pattern * occurrences) option -> clause -> unit := "ltac2" "tac_native". +Ltac2 @ external eval_red : constr -> constr := "ltac2" "eval_red". +Ltac2 @ external eval_hnf : constr -> constr := "ltac2" "eval_hnf". +Ltac2 @ external eval_red : constr -> constr := "ltac2" "eval_red". +Ltac2 @ external eval_simpl : red_flags -> (pattern * occurrences) option -> constr -> constr := "ltac2" "eval_simpl". +Ltac2 @ external eval_cbv : red_flags -> constr -> constr := "ltac2" "eval_cbv". +Ltac2 @ external eval_cbn : red_flags -> constr -> constr := "ltac2" "eval_cbn". +Ltac2 @ external eval_lazy : red_flags -> constr -> constr := "ltac2" "eval_lazy". +Ltac2 @ external eval_unfold : (reference * occurrences) list -> constr -> constr := "ltac2" "eval_unfold". +Ltac2 @ external eval_fold : constr list -> constr -> constr := "ltac2" "eval_fold". +Ltac2 @ external eval_pattern : (constr * occurrences) list -> constr -> constr := "ltac2" "eval_pattern". +Ltac2 @ external eval_vm : (pattern * occurrences) option -> constr -> constr := "ltac2" "eval_vm". +Ltac2 @ external eval_native : (pattern * occurrences) option -> constr -> constr := "ltac2" "eval_native". + Ltac2 @ external rewrite : evar_flag -> rewriting list -> clause -> (unit -> unit) option -> unit := "ltac2" "tac_rewrite". Ltac2 @ external reflexivity : unit -> unit := "ltac2" "tac_reflexivity". -- cgit v1.2.3 From 84047666ce13f1eec440d38d9784ae125612507c Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 31 Aug 2017 18:09:05 +0200 Subject: Fix the type of the Constructor constructor. --- theories/Constr.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/Constr.v b/theories/Constr.v index bb02d94531..801192d628 100644 --- a/theories/Constr.v +++ b/theories/Constr.v @@ -31,7 +31,7 @@ Ltac2 Type kind := [ | App (constr, constr array) | Constant (constant, instance) | Ind (inductive, instance) -| Constructor (inductive, instance) +| Constructor (constructor, instance) | Case (constr, constr, constr array) | Fix (int array, int, ident option array, constr array, constr array) | CoFix (int, ident option array, constr array, constr array) -- cgit v1.2.3 From 40edf6a111ae2b9f0a230c2617b3e86e8fbfa6dd Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 31 Aug 2017 18:15:09 +0200 Subject: Fix coq/ltac2#3: Constructor expects n arguments should name which constructor it is. --- src/tac2intern.ml | 17 ++++++++++------- src/tac2intern.mli | 2 +- 2 files changed, 11 insertions(+), 8 deletions(-) diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 051b3af5c7..88824386d9 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -206,9 +206,10 @@ let loc_of_patexpr = function | CPatVar (loc, _) -> Option.default dummy_loc loc | CPatRef (loc, _, _) -> loc -let error_nargs_mismatch loc nargs nfound = - user_err ~loc (str "Constructor expects " ++ int nargs ++ - str " arguments, but is applied to " ++ int nfound ++ +let error_nargs_mismatch loc kn nargs nfound = + let cstr = Tac2env.shortest_qualid_of_constructor kn in + user_err ~loc (str "Constructor " ++ pr_qualid cstr ++ str " expects " ++ + int nargs ++ str " arguments, but is applied to " ++ int nfound ++ str " arguments") let error_nparams_mismatch loc nargs nfound = @@ -924,8 +925,10 @@ and intern_case env loc e pl = let ids = List.map get_id args in let nids = List.length ids in let nargs = List.length arity in - let () = - if not (Int.equal nids nargs) then error_nargs_mismatch loc nargs nids + let () = match knc with + | Tuple n -> assert (n == nids) + | Other knc -> + if not (Int.equal nids nargs) then error_nargs_mismatch loc knc nargs nids in let fold env id tpe = (** Instantiate all arguments *) @@ -993,7 +996,7 @@ and intern_case env loc e pl = let nids = List.length ids in let nargs = List.length data.cdata_args in let () = - if not (Int.equal nids nargs) then error_nargs_mismatch loc nargs nids + if not (Int.equal nids nargs) then error_nargs_mismatch loc knc nargs nids in let fold env id tpe = (** Instantiate all arguments *) @@ -1033,7 +1036,7 @@ and intern_constructor env loc kn args = match kn with | None -> (GTacOpn (kn, args), ans) else - error_nargs_mismatch loc nargs (List.length args) + error_nargs_mismatch loc kn nargs (List.length args) | Tuple n -> assert (Int.equal n (List.length args)); let types = List.init n (fun i -> GTypVar (fresh_id env)) in diff --git a/src/tac2intern.mli b/src/tac2intern.mli index dac074a0eb..8c997b741e 100644 --- a/src/tac2intern.mli +++ b/src/tac2intern.mli @@ -41,5 +41,5 @@ val globalize : Id.Set.t -> raw_tacexpr -> raw_tacexpr (** Errors *) -val error_nargs_mismatch : Loc.t -> int -> int -> 'a +val error_nargs_mismatch : Loc.t -> ltac_constructor -> int -> int -> 'a val error_nparams_mismatch : Loc.t -> int -> int -> 'a -- cgit v1.2.3 From edc4126a37d7ea8f99142b706c9e6b6eb806443e Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Thu, 31 Aug 2017 12:32:51 -0400 Subject: Require Ltac2.Fresh in Ltac2.Ltac2 --- theories/Ltac2.v | 1 + 1 file changed, 1 insertion(+) diff --git a/theories/Ltac2.v b/theories/Ltac2.v index 07229797da..996236325c 100644 --- a/theories/Ltac2.v +++ b/theories/Ltac2.v @@ -15,6 +15,7 @@ Require Ltac2.Array. Require Ltac2.Message. Require Ltac2.Constr. Require Ltac2.Control. +Require Ltac2.Fresh. Require Ltac2.Pattern. Require Ltac2.Std. Require Ltac2.Notations. -- cgit v1.2.3 From 7efbf5add76d640b5083110a5163bb8c1b98dabd Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 31 Aug 2017 18:55:35 +0200 Subject: Fix coq/ltac2#10: Antiquotation syntax breaks when backtracking across `Require`. --- src/g_ltac2.ml4 | 2 ++ src/tac2entries.ml | 12 ++++++++++++ src/tac2entries.mli | 4 ++++ 3 files changed, 18 insertions(+) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 4e62fab715..bfd2ab1a11 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -650,6 +650,7 @@ END (** Extension of constr syntax *) +let () = Hook.set Tac2entries.register_constr_quotations begin fun () -> GEXTEND Gram Pcoq.Constr.operconstr: LEVEL "0" [ [ IDENT "ltac2"; ":"; "("; tac = tac2expr; ")" -> @@ -666,6 +667,7 @@ GEXTEND Gram ] ] ; END +end let pr_ltac2entry _ = mt () (** FIXME *) let pr_ltac2expr _ = mt () (** FIXME *) diff --git a/src/tac2entries.ml b/src/tac2entries.ml index da7c07c134..2895fda060 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -819,10 +819,22 @@ let def_unit = { let t_list = coq_def "list" +let (f_register_constr_quotations, register_constr_quotations) = Hook.make () + +let perform_constr_quotations (_, ()) = Hook.get f_register_constr_quotations () + +(** Dummy object that register global rules when Require is called *) +let inTac2ConstrQuotations : unit -> obj = + declare_object {(default_object "TAC2-CONSTR-QUOT") with + cache_function = perform_constr_quotations; + load_function = fun _ -> perform_constr_quotations; + } + let _ = Mltop.declare_cache_obj begin fun () -> ignore (Lib.add_leaf (Id.of_string "unit") (inTypDef def_unit)); register_prim_alg "list" 1 [ ("[]", []); ("::", [GTypVar 0; GTypRef (Other t_list, [GTypVar 0])]); ]; + Lib.add_anonymous_leaf (inTac2ConstrQuotations ()); end "ltac2_plugin" diff --git a/src/tac2entries.mli b/src/tac2entries.mli index acb99a34b1..7ed45e62e5 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -72,3 +72,7 @@ val q_occurrences : occurrences Pcoq.Gram.entry val q_reference : Libnames.reference or_anti Pcoq.Gram.entry val q_strategy_flag : strategy_flag Pcoq.Gram.entry end + +(** {5 Hooks} *) + +val register_constr_quotations : (unit -> unit) Hook.t -- cgit v1.2.3 From e89c5c3de0f00de2732f385087a3461b4e6f3a84 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 31 Aug 2017 19:36:50 +0200 Subject: Expand the primitive functions on terms. --- src/tac2core.ml | 131 +++++++++++++++++++++++++++++++++++++++++++++++++++--- src/tac2ffi.ml | 1 + src/tac2ffi.mli | 1 + theories/Constr.v | 15 ++++++- 4 files changed, 141 insertions(+), 7 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 95fd29ec33..1f500352cf 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -61,15 +61,32 @@ let of_name c = match c with | Anonymous -> Value.of_option Value.of_ident None | Name id -> Value.of_option Value.of_ident (Some id) -let of_instance sigma u = - let u = Univ.Instance.to_array (EConstr.EInstance.kind sigma u) in +let to_name c = match Value.to_option Value.to_ident c with +| None -> Anonymous +| Some id -> Name id + +let of_instance u = + let u = Univ.Instance.to_array (EConstr.Unsafe.to_instance u) in Value.of_array (fun v -> Value.of_ext Value.val_univ v) u +let to_instance u = + let u = Value.to_array (fun v -> Value.to_ext Value.val_univ v) u in + EConstr.EInstance.make (Univ.Instance.of_array u) + let of_rec_declaration (nas, ts, cs) = (Value.of_array of_name nas, Value.of_array Value.of_constr ts, Value.of_array Value.of_constr cs) +let to_rec_declaration (nas, ts, cs) = + (Value.to_array to_name nas, + Value.to_array Value.to_constr ts, + Value.to_array Value.to_constr cs) + +let of_result f = function +| Inl c -> ValBlk (0, [|f c|]) +| Inr e -> ValBlk (1, [|Value.of_exn e|]) + (** Stdlib exceptions *) let err_notfocussed = @@ -335,20 +352,21 @@ let () = define1 "constr_kind" begin fun c -> | Const (cst, u) -> ValBlk (10, [| Value.of_constant cst; - of_instance sigma u; + of_instance u; |]) | Ind (ind, u) -> ValBlk (11, [| Value.of_ext Value.val_inductive ind; - of_instance sigma u; + of_instance u; |]) | Construct (cstr, u) -> ValBlk (12, [| Value.of_ext Value.val_constructor cstr; - of_instance sigma u; + of_instance u; |]) - | Case (_, c, t, bl) -> + | Case (ci, c, t, bl) -> ValBlk (13, [| + Value.of_ext Value.val_case ci; Value.of_constr c; Value.of_constr t; Value.of_array Value.of_constr bl; @@ -378,6 +396,99 @@ let () = define1 "constr_kind" begin fun c -> end end +let () = define1 "constr_make" begin fun knd -> + let open Constr in + let c = match knd with + | ValBlk (0, [|n|]) -> + let n = Value.to_int n in + EConstr.mkRel n + | ValBlk (1, [|id|]) -> + let id = Value.to_ident id in + EConstr.mkVar id + | ValBlk (2, [|n|]) -> + let n = Value.to_int n in + EConstr.mkMeta n + | ValBlk (3, [|evk; args|]) -> + let evk = Evar.unsafe_of_int (Value.to_int evk) in + let args = Value.to_array Value.to_constr args in + EConstr.mkEvar (evk, args) + | ValBlk (4, [|s|]) -> + let s = Value.to_ext Value.val_sort s in + EConstr.mkSort (EConstr.Unsafe.to_sorts s) + | ValBlk (5, [|c; k; t|]) -> + let c = Value.to_constr c in + let k = Value.to_ext Value.val_cast k in + let t = Value.to_constr t in + EConstr.mkCast (c, k, t) + | ValBlk (6, [|na; t; u|]) -> + let na = to_name na in + let t = Value.to_constr t in + let u = Value.to_constr u in + EConstr.mkProd (na, t, u) + | ValBlk (7, [|na; t; c|]) -> + let na = to_name na in + let t = Value.to_constr t in + let u = Value.to_constr c in + EConstr.mkLambda (na, t, u) + | ValBlk (8, [|na; b; t; c|]) -> + let na = to_name na in + let b = Value.to_constr b in + let t = Value.to_constr t in + let c = Value.to_constr c in + EConstr.mkLetIn (na, b, t, c) + | ValBlk (9, [|c; cl|]) -> + let c = Value.to_constr c in + let cl = Value.to_array Value.to_constr cl in + EConstr.mkApp (c, cl) + | ValBlk (10, [|cst; u|]) -> + let cst = Value.to_constant cst in + let u = to_instance u in + EConstr.mkConstU (cst, u) + | ValBlk (11, [|ind; u|]) -> + let ind = Value.to_ext Value.val_inductive ind in + let u = to_instance u in + EConstr.mkIndU (ind, u) + | ValBlk (12, [|cstr; u|]) -> + let cstr = Value.to_ext Value.val_constructor cstr in + let u = to_instance u in + EConstr.mkConstructU (cstr, u) + | ValBlk (13, [|ci; c; t; bl|]) -> + let ci = Value.to_ext Value.val_case ci in + let c = Value.to_constr c in + let t = Value.to_constr t in + let bl = Value.to_array Value.to_constr bl in + EConstr.mkCase (ci, c, t, bl) + | ValBlk (14, [|recs; i; nas; ts; cs|]) -> + let recs = Value.to_array Value.to_int recs in + let i = Value.to_int i in + let def = to_rec_declaration (nas, ts, cs) in + EConstr.mkFix ((recs, i), def) + | ValBlk (15, [|i; nas; ts; cs|]) -> + let i = Value.to_int i in + let def = to_rec_declaration (nas, ts, cs) in + EConstr.mkCoFix (i, def) + | ValBlk (16, [|p; c|]) -> + let p = Value.to_ext Value.val_projection p in + let c = Value.to_constr c in + EConstr.mkProj (p, c) + | _ -> assert false + in + return (Value.of_constr c) +end + +let () = define1 "constr_check" begin fun c -> + let c = Value.to_constr c in + pf_apply begin fun env sigma -> + try + let (sigma, _) = Typing.type_of env sigma c in + Proofview.Unsafe.tclEVARS sigma >>= fun () -> + return (of_result Value.of_constr (Inl c)) + with e when CErrors.noncritical e -> + let e = CErrors.push e in + return (of_result Value.of_constr (Inr e)) + end +end + let () = define3 "constr_substnl" begin fun subst k c -> let subst = Value.to_list Value.to_constr subst in let k = Value.to_int k in @@ -386,6 +497,14 @@ let () = define3 "constr_substnl" begin fun subst k c -> return (Value.of_constr ans) end +let () = define3 "constr_closenl" begin fun ids k c -> + let ids = Value.to_list Value.to_ident ids in + let k = Value.to_int k in + let c = Value.to_constr c in + let ans = EConstr.Vars.substn_vars k ids c in + return (Value.of_constr ans) +end + (** Patterns *) let () = define2 "pattern_matches" begin fun pat c -> diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index 4d84da14ce..a1f9debdcb 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -26,6 +26,7 @@ let val_inductive = Val.create "inductive" let val_constant = Val.create "constant" let val_constructor = Val.create "constructor" let val_projection = Val.create "projection" +let val_case = Val.create "case" let val_univ = Val.create "universe" let val_kont : (Exninfo.iexn -> valexpr Proofview.tactic) Val.tag = Val.create "kont" diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index cf1d7e81a1..e836319349 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -80,6 +80,7 @@ val val_inductive : inductive Val.tag val val_constant : Constant.t Val.tag val val_constructor : constructor Val.tag val val_projection : Projection.t Val.tag +val val_case : Constr.case_info Val.tag val val_univ : Univ.universe_level Val.tag val val_kont : (Exninfo.iexn -> valexpr Proofview.tactic) Val.tag val val_free : Id.Set.t Val.tag diff --git a/theories/Constr.v b/theories/Constr.v index 801192d628..3e67a486cf 100644 --- a/theories/Constr.v +++ b/theories/Constr.v @@ -18,6 +18,8 @@ Module Unsafe. (** Low-level access to kernel terms. Use with care! *) +Ltac2 Type case. + Ltac2 Type kind := [ | Rel (int) | Var (ident) @@ -32,7 +34,7 @@ Ltac2 Type kind := [ | Constant (constant, instance) | Ind (inductive, instance) | Constructor (constructor, instance) -| Case (constr, constr, constr array) +| Case (case, constr, constr, constr array) | Fix (int array, int, ident option array, constr array, constr array) | CoFix (int, ident option array, constr array, constr array) | Proj (projection, constr) @@ -40,8 +42,19 @@ Ltac2 Type kind := [ Ltac2 @ external kind : constr -> kind := "ltac2" "constr_kind". +Ltac2 @ external make : kind -> constr := "ltac2" "constr_make". + +Ltac2 @ external check : constr -> constr result := "ltac2" "constr_check". +(** Checks that a constr generated by unsafe means is indeed safe in the + current environment, and returns it, or the error otherwise. Panics if + not focussed. *) + Ltac2 @ external substnl : constr list -> int -> constr -> constr := "ltac2" "constr_substnl". (** [substnl [r₁;...;rₙ] k c] substitutes in parallel [Rel(k+1); ...; Rel(k+n)] with [r₁;...;rₙ] in [c]. *) +Ltac2 @ external closenl : ident list -> int -> constr -> constr := "ltac2" "constr_closenl". +(** [closenl [x₁;...;xₙ] k c] abstracts over variables [x₁;...;xₙ] and replaces them with + [Rel(k); ...; Rel(k+n-1)] in [c]. If two names are identical, the one of least index is kept. *) + End Unsafe. -- cgit v1.2.3 From 72e3d2e563e08627559065ff0289403591d99682 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 31 Aug 2017 23:49:21 +0200 Subject: Properly handling internal errors from Coq. --- src/tac2ffi.ml | 14 +++++++++++--- tests/errors.v | 12 ++++++++++++ theories/Init.v | 6 ++++++ 3 files changed, 29 insertions(+), 3 deletions(-) create mode 100644 tests/errors.v diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index a1f9debdcb..dd20de5ef5 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -92,14 +92,22 @@ let to_ident c = to_ext val_ident c let of_pattern c = of_ext val_pattern c let to_pattern c = to_ext val_pattern c +let internal_err = + let open Names in + KerName.make2 Tac2env.coq_prefix (Label.of_id (Id.of_string "Internal")) + (** FIXME: handle backtrace in Ltac2 exceptions *) let of_exn c = match fst c with | LtacError (kn, c) -> ValOpn (kn, c) -| _ -> of_ext val_exn c +| _ -> ValOpn (internal_err, [|of_ext val_exn c|]) let to_exn c = match c with -| ValOpn (kn, c) -> (LtacError (kn, c), Exninfo.null) -| _ -> to_ext val_exn c +| ValOpn (kn, c) -> + if Names.KerName.equal kn internal_err then + to_ext val_exn c.(0) + else + (LtacError (kn, c), Exninfo.null) +| _ -> assert false let of_option f = function | None -> ValInt 0 diff --git a/tests/errors.v b/tests/errors.v new file mode 100644 index 0000000000..e7beff3420 --- /dev/null +++ b/tests/errors.v @@ -0,0 +1,12 @@ +Require Import Ltac2.Ltac2. + +Goal True. +Proof. +let x := Control.plus + (fun () => let _ := constr:(nat -> 0) in 0) + (fun e => match e with Not_found => 1 | _ => 2 end) in +match Int.equal x 2 with +| true => () +| false => Control.throw Tactic_failure +end. +Abort. diff --git a/theories/Init.v b/theories/Init.v index 1591747eb4..baaf5956b2 100644 --- a/theories/Init.v +++ b/theories/Init.v @@ -46,6 +46,12 @@ Ltac2 Type 'a result := [ Val ('a) | Err (exn) ]. (** Pervasive exceptions *) +Ltac2 Type err. +(** Coq internal errors. Cannot be constructed, merely passed around. *) + +Ltac2 Type exn ::= [ Internal (err) ]. +(** Wrapper around the errors raised by Coq implementation. *) + Ltac2 Type exn ::= [ Out_of_bounds ]. (** Used for bound checking, e.g. with String and Array. *) -- cgit v1.2.3 From 2a0a48834f0b90319e56ae9f4a172fe6855583c0 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 1 Sep 2017 00:19:33 +0200 Subject: Passing an optional message to Tactic_failure. --- tests/errors.v | 2 +- theories/Init.v | 2 +- theories/Notations.v | 8 ++++---- 3 files changed, 6 insertions(+), 6 deletions(-) diff --git a/tests/errors.v b/tests/errors.v index e7beff3420..c677f6af5d 100644 --- a/tests/errors.v +++ b/tests/errors.v @@ -7,6 +7,6 @@ let x := Control.plus (fun e => match e with Not_found => 1 | _ => 2 end) in match Int.equal x 2 with | true => () -| false => Control.throw Tactic_failure +| false => Control.throw (Tactic_failure None) end. Abort. diff --git a/theories/Init.v b/theories/Init.v index baaf5956b2..04394e2c5d 100644 --- a/theories/Init.v +++ b/theories/Init.v @@ -67,5 +67,5 @@ Ltac2 Type exn ::= [ Not_found ]. Ltac2 Type exn ::= [ Match_failure ]. (** Used to signal a pattern didn't match a term. *) -Ltac2 Type exn ::= [ Tactic_failure ]. +Ltac2 Type exn ::= [ Tactic_failure (message option) ]. (** Generic error for tactic failure. *) diff --git a/theories/Notations.v b/theories/Notations.v index 46c0e5e79f..411367eab1 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -27,7 +27,7 @@ match Control.case t with Control.plus (fun _ => s x) (fun e => s (k e)) end. -Ltac2 fail0 (_ : unit) := Control.enter (fun _ => Control.zero Tactic_failure). +Ltac2 fail0 (_ : unit) := Control.enter (fun _ => Control.zero (Tactic_failure None)). Ltac2 Notation fail := fail0 (). @@ -69,7 +69,7 @@ Ltac2 Notation progress := progress0. Ltac2 rec first0 tacs := match tacs with -| [] => Control.zero Tactic_failure +| [] => Control.zero (Tactic_failure None) | tac :: tacs => Control.enter (fun _ => orelse tac (fun _ => first0 tacs)) end. @@ -77,12 +77,12 @@ Ltac2 Notation "first" "[" tacs(list0(thunk(tactic(6)), "|")) "]" := first0 tacs Ltac2 complete tac := let ans := tac () in - Control.enter (fun () => Control.zero Tactic_failure); + Control.enter (fun () => Control.zero (Tactic_failure None)); ans. Ltac2 rec solve0 tacs := match tacs with -| [] => Control.zero Tactic_failure +| [] => Control.zero (Tactic_failure None) | tac :: tacs => Control.enter (fun _ => orelse (fun _ => complete tac) (fun _ => solve0 tacs)) end. -- cgit v1.2.3 From 1f7a2ea0e0513620bb946c10923d38f845061afa Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 1 Sep 2017 00:27:35 +0200 Subject: Ensuring the Ltac definitions have lowercase names. --- src/tac2entries.ml | 8 +++++++- 1 file changed, 7 insertions(+), 1 deletion(-) diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 2895fda060..efdc995a69 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -307,6 +307,10 @@ let inline_rec_tactic tactics = in List.map map tactics +let check_lowercase (loc, id) = + if Tac2env.is_constructor (Libnames.qualid_of_ident id) then + user_err ?loc (str "The identifier " ++ Id.print id ++ str " must be lowercase") + let register_ltac ?(local = false) ?(mut = false) isrec tactics = let map ((loc, na), e) = let id = match na with @@ -314,6 +318,7 @@ let register_ltac ?(local = false) ?(mut = false) isrec tactics = user_err ?loc (str "Tactic definition must have a name") | Name id -> id in + let () = check_lowercase (loc, id) in ((loc, id), e) in let tactics = List.map map tactics in @@ -648,8 +653,9 @@ let inTac2Abbreviation : abbreviation -> obj = classify_function = classify_abbreviation} let register_notation ?(local = false) tkn lev body = match tkn, lev with -| [SexprRec (_, (_, Some id), [])], None -> +| [SexprRec (_, (loc, Some id), [])], None -> (** Tactic abbreviation *) + let () = check_lowercase (loc, id) in let body = Tac2intern.globalize Id.Set.empty body in let abbr = { abbr_body = body } in ignore (Lib.add_leaf id (inTac2Abbreviation abbr)) -- cgit v1.2.3 From 0efde84daaa6b032e40a920a36793181724de87a Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 2 Sep 2017 17:38:31 +0200 Subject: Fix coq/ltac2#15: Ltac2 scope cannot be disabled after it is required. Instead of setting globally the option, we add a hook to set it in the init object of the plugin. --- src/tac2entries.ml | 20 ++++++++++++++------ theories/Init.v | 2 -- 2 files changed, 14 insertions(+), 8 deletions(-) diff --git a/src/tac2entries.ml b/src/tac2entries.ml index efdc995a69..a503a0ab4c 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -827,13 +827,21 @@ let t_list = coq_def "list" let (f_register_constr_quotations, register_constr_quotations) = Hook.make () -let perform_constr_quotations (_, ()) = Hook.get f_register_constr_quotations () +let cache_ltac2_init (_, ()) = + Hook.get f_register_constr_quotations () + +let load_ltac2_init _ (_, ()) = + Hook.get f_register_constr_quotations () + +let open_ltac2_init _ (_, ()) = + Goptions.set_string_option_value_gen None ["Default"; "Proof"; "Mode"] "Ltac2" (** Dummy object that register global rules when Require is called *) -let inTac2ConstrQuotations : unit -> obj = - declare_object {(default_object "TAC2-CONSTR-QUOT") with - cache_function = perform_constr_quotations; - load_function = fun _ -> perform_constr_quotations; +let inTac2Init : unit -> obj = + declare_object {(default_object "TAC2-INIT") with + cache_function = cache_ltac2_init; + load_function = load_ltac2_init; + open_function = open_ltac2_init; } let _ = Mltop.declare_cache_obj begin fun () -> @@ -842,5 +850,5 @@ let _ = Mltop.declare_cache_obj begin fun () -> ("[]", []); ("::", [GTypVar 0; GTypRef (Other t_list, [GTypVar 0])]); ]; - Lib.add_anonymous_leaf (inTac2ConstrQuotations ()); + Lib.add_anonymous_leaf (inTac2Init ()); end "ltac2_plugin" diff --git a/theories/Init.v b/theories/Init.v index 04394e2c5d..16e7d7a6f9 100644 --- a/theories/Init.v +++ b/theories/Init.v @@ -8,8 +8,6 @@ Declare ML Module "ltac2_plugin". -Global Set Default Proof Mode "Ltac2". - (** Primitive types *) Ltac2 Type int. -- cgit v1.2.3 From f84c0b96f11d5d1f130d36c0c04597ddeca6001f Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 2 Sep 2017 18:19:34 +0200 Subject: Fix coq/ltac2#12: Error should name which match cases are unhandled. --- src/tac2intern.ml | 11 +++++++---- src/tac2print.ml | 16 ++++++++++++---- src/tac2print.mli | 1 + 3 files changed, 20 insertions(+), 8 deletions(-) diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 88824386d9..b9e77f3cf5 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -953,12 +953,15 @@ and intern_case env loc e pl = intern_branch rem in let () = intern_branch pl in - let map = function - | None -> user_err ~loc (str "TODO: Unhandled match case") (** FIXME *) + let map n is_const = function + | None -> + let kn = match kn with Other kn -> kn | _ -> assert false in + let cstr = pr_internal_constructor kn n is_const in + user_err ~loc (str "Unhandled match case for constructor " ++ cstr) | Some x -> x in - let const = Array.map map const in - let nonconst = Array.map map nonconst in + let const = Array.mapi (fun i o -> map i true o) const in + let nonconst = Array.mapi (fun i o -> map i false o) nonconst in let ce = GTacCse (e', kn, const, nonconst) in (ce, ret) | PKind_open kn -> diff --git a/src/tac2print.ml b/src/tac2print.ml index 939f44aeaf..75ad2082d4 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -125,6 +125,15 @@ let find_constructor n empty def = in find n def +let pr_internal_constructor tpe n is_const = + let data = match Tac2env.interp_type tpe with + | (_, GTydAlg data) -> data + | _ -> assert false + in + let id = find_constructor n is_const data.galg_constructors in + let kn = change_kn_label tpe id in + pr_constructor kn + let order_branches cbr nbr def = let rec order cidx nidx def = match def with | [] -> [] @@ -179,18 +188,17 @@ let pr_glbexpr_gen lvl c = paren (prlist_with_sep (fun () -> str "," ++ spc ()) (pr_glbexpr E1) cl) | GTacCst (Other tpe, n, cl) -> begin match Tac2env.interp_type tpe with - | _, GTydAlg { galg_constructors = def } -> + | _, GTydAlg def -> let paren = match lvl with | E0 -> paren | E1 | E2 | E3 | E4 | E5 -> fun x -> x in - let id = find_constructor n (List.is_empty cl) def in - let kn = change_kn_label tpe id in + let cstr = pr_internal_constructor tpe n (List.is_empty cl) in let cl = match cl with | [] -> mt () | _ -> spc () ++ pr_sequence (pr_glbexpr E0) cl in - paren (hov 2 (pr_constructor kn ++ cl)) + paren (hov 2 (cstr ++ cl)) | _, GTydRec def -> let args = List.combine def cl in let pr_arg ((id, _, _), arg) = diff --git a/src/tac2print.mli b/src/tac2print.mli index 0a04af2ff0..737e813ed3 100644 --- a/src/tac2print.mli +++ b/src/tac2print.mli @@ -24,6 +24,7 @@ val pr_glbtype : ('a -> string) -> 'a glb_typexpr -> Pp.t (** {5 Printing expressions} *) val pr_constructor : ltac_constructor -> Pp.t +val pr_internal_constructor : type_constant -> int -> bool -> Pp.t val pr_projection : ltac_projection -> Pp.t val pr_glbexpr_gen : exp_level -> glb_tacexpr -> Pp.t val pr_glbexpr : glb_tacexpr -> Pp.t -- cgit v1.2.3 From 4d5e3f3f00cb1848861b938ba1a57c33800d71a6 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 3 Sep 2017 01:59:34 +0200 Subject: Fix coq/ltac2#16: Passing Ltac2 variables to Ltac1 via $ results in anomalies. --- src/tac2core.ml | 3 +++ src/tac2intern.ml | 3 +++ src/tac2intern.mli | 4 ++++ tests/compat.v | 4 ++++ 4 files changed, 14 insertions(+) diff --git a/src/tac2core.ml b/src/tac2core.ml index 1f500352cf..81388af0ef 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -873,6 +873,9 @@ let () = let () = let intern ist tac = + (** Prevent inner calls to Ltac2 values *) + let extra = Tac2intern.drop_ltac2_env ist.Genintern.extra in + let ist = { ist with Genintern.extra } in let _, tac = Genintern.intern Ltac_plugin.Tacarg.wit_tactic ist tac in tac in diff --git a/src/tac2intern.ml b/src/tac2intern.ml index b9e77f3cf5..490436422d 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -169,6 +169,9 @@ let env_name env = let ltac2_env : environment Genintern.Store.field = Genintern.Store.field () +let drop_ltac2_env store = + Genintern.Store.remove store ltac2_env + let fresh_id env = UF.fresh env.env_cst let get_alias (loc, id) env = diff --git a/src/tac2intern.mli b/src/tac2intern.mli index 8c997b741e..95199d449d 100644 --- a/src/tac2intern.mli +++ b/src/tac2intern.mli @@ -43,3 +43,7 @@ val globalize : Id.Set.t -> raw_tacexpr -> raw_tacexpr val error_nargs_mismatch : Loc.t -> ltac_constructor -> int -> int -> 'a val error_nparams_mismatch : Loc.t -> int -> int -> 'a + +(** Misc *) + +val drop_ltac2_env : Genintern.Store.t -> Genintern.Store.t diff --git a/tests/compat.v b/tests/compat.v index 44421349da..f4e849c5de 100644 --- a/tests/compat.v +++ b/tests/compat.v @@ -25,6 +25,9 @@ Abort. (** Variables do not cross the compatibility layer boundary. *) Fail Ltac2 bar nay := ltac1:(discriminate nay). +Fail Ltac2 pose1 (v : constr) := + ltac1:(pose $v). + (** Test calls to Ltac2 from Ltac1 *) Set Default Proof Mode "Classic". @@ -54,3 +57,4 @@ Fail mytac ltac2:(fail). let t := idtac; ltac2:(fail) in mytac t. constructor. Qed. + -- cgit v1.2.3 From 83a92df4e2e94bfc33354cf26627329d4a2e0610 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 3 Sep 2017 17:39:16 +0200 Subject: Allowing complex types in ML objects. --- src/tac2core.ml | 29 +++++++++++++---------------- src/tac2env.ml | 5 +++-- src/tac2env.mli | 5 +++-- src/tac2intern.ml | 13 ++++++++++--- 4 files changed, 29 insertions(+), 23 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 81388af0ef..b8bec37d04 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -775,9 +775,11 @@ let to_lvar ist = let lfun = Tac2interp.set_env ist Id.Map.empty in { empty_lvar with Glob_term.ltac_genargs = lfun } -let intern_constr ist c = +let gtypref kn = GTypRef (Other kn, []) + +let intern_constr self ist c = let (_, (c, _)) = Genintern.intern Stdarg.wit_constr ist c in - c + (c, gtypref t_constr) let interp_constr flags ist c = let open Pretyping in @@ -796,7 +798,6 @@ let () = let interp ist c = interp_constr (constr_flags ()) ist c in let print env c = str "constr:(" ++ Printer.pr_lglob_constr_env env c ++ str ")" in let obj = { - ml_type = t_constr; ml_intern = intern; ml_subst = Detyping.subst_glob_constr; ml_interp = interp; @@ -809,7 +810,6 @@ let () = let interp ist c = interp_constr (open_constr_no_classes_flags ()) ist c in let print env c = str "open_constr:(" ++ Printer.pr_lglob_constr_env env c ++ str ")" in let obj = { - ml_type = t_constr; ml_intern = intern; ml_subst = Detyping.subst_glob_constr; ml_interp = interp; @@ -821,8 +821,7 @@ let () = let interp _ id = return (ValExt (Value.val_ident, id)) in let print _ id = str "ident:(" ++ Id.print id ++ str ")" in let obj = { - ml_type = t_ident; - ml_intern = (fun _ id -> id); + ml_intern = (fun _ _ id -> id, gtypref t_ident); ml_interp = interp; ml_subst = (fun _ id -> id); ml_print = print; @@ -830,14 +829,13 @@ let () = define_ml_object Tac2env.wit_ident obj let () = - let intern ist c = + let intern self ist c = let _, pat = Constrintern.intern_constr_pattern ist.Genintern.genv ~as_type:false c in - pat + pat, gtypref t_pattern in let print env pat = str "pattern:(" ++ Printer.pr_lconstr_pattern_env env Evd.empty pat ++ str ")" in let interp _ c = return (ValExt (Value.val_pattern, c)) in let obj = { - ml_type = t_pattern; ml_intern = intern; ml_interp = interp; ml_subst = Patternops.subst_pattern; @@ -846,15 +844,16 @@ let () = define_ml_object Tac2env.wit_pattern obj let () = - let intern ist qid = match qid with - | Libnames.Ident (_, id) -> Globnames.VarRef id + let intern self ist qid = match qid with + | Libnames.Ident (_, id) -> + Globnames.VarRef id, gtypref t_reference | Libnames.Qualid (loc, qid) -> let gr = try Nametab.locate qid with Not_found -> Nametab.error_global_not_found ?loc qid in - gr + gr, gtypref t_reference in let subst s c = Globnames.subst_global_reference s c in let interp _ gr = return (Value.of_reference gr) in @@ -863,7 +862,6 @@ let () = | r -> str "reference:(" ++ Printer.pr_global r ++ str ")" in let obj = { - ml_type = t_reference; ml_intern = intern; ml_subst = subst; ml_interp = interp; @@ -872,12 +870,12 @@ let () = define_ml_object Tac2env.wit_reference obj let () = - let intern ist tac = + let intern self ist tac = (** Prevent inner calls to Ltac2 values *) let extra = Tac2intern.drop_ltac2_env ist.Genintern.extra in let ist = { ist with Genintern.extra } in let _, tac = Genintern.intern Ltac_plugin.Tacarg.wit_tactic ist tac in - tac + tac, gtypref t_unit in let interp _ tac = (** FUCK YOU API *) @@ -889,7 +887,6 @@ let () = str "ltac1:(" ++ Ltac_plugin.Pptactic.pr_glob_tactic (Obj.magic env) tac ++ str ")" in let obj = { - ml_type = t_unit; ml_intern = intern; ml_subst = subst; ml_interp = interp; diff --git a/src/tac2env.ml b/src/tac2env.ml index 489113c031..56fd55ee84 100644 --- a/src/tac2env.ml +++ b/src/tac2env.ml @@ -251,9 +251,10 @@ let shortest_qualid_of_projection kn = let sp = KNmap.find kn tab.tab_proj_rev in KnTab.shortest_qualid Id.Set.empty sp tab.tab_proj +type ('a, 'b, 'r) intern_fun = Genintern.glob_sign -> 'a -> 'b * 'r glb_typexpr + type ('a, 'b) ml_object = { - ml_type : type_constant; - ml_intern : Genintern.glob_sign -> 'a -> 'b; + ml_intern : 'r. (raw_tacexpr, glb_tacexpr, 'r) intern_fun -> ('a, 'b, 'r) intern_fun; ml_subst : Mod_subst.substitution -> 'b -> 'b; ml_interp : environment -> 'b -> valexpr Proofview.tactic; ml_print : Environ.env -> 'b -> Pp.t; diff --git a/src/tac2env.mli b/src/tac2env.mli index 0ef62d67ed..15664db756 100644 --- a/src/tac2env.mli +++ b/src/tac2env.mli @@ -105,9 +105,10 @@ val interp_primitive : ml_tactic_name -> ml_tactic (** {5 ML primitive types} *) +type ('a, 'b, 'r) intern_fun = Genintern.glob_sign -> 'a -> 'b * 'r glb_typexpr + type ('a, 'b) ml_object = { - ml_type : type_constant; - ml_intern : Genintern.glob_sign -> 'a -> 'b; + ml_intern : 'r. (raw_tacexpr, glb_tacexpr, 'r) intern_fun -> ('a, 'b, 'r) intern_fun; ml_subst : Mod_subst.substitution -> 'b -> 'b; ml_interp : environment -> 'b -> valexpr Proofview.tactic; ml_print : Environ.env -> 'b -> Pp.t; diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 490436422d..5d2fc2b47b 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -775,14 +775,21 @@ let rec intern_rec env = function (GTacSet (pinfo.pdata_type, e, pinfo.pdata_indx, r), GTypRef (Tuple 0, [])) | CTacExt (loc, tag, arg) -> let open Genintern in - let tpe = interp_ml_object tag in + let self ist e = + let env = match Store.get ist.extra ltac2_env with + | None -> empty_env () + | Some env -> env + in + intern_rec env e + in + let obj = interp_ml_object tag in (** External objects do not have access to the named context because this is not stable by dynamic semantics. *) let genv = Global.env_of_context Environ.empty_named_context_val in let ist = empty_glob_sign genv in let ist = { ist with extra = Store.set ist.extra ltac2_env env } in - let arg = Flags.with_option Ltac_plugin.Tacintern.strict_check (fun () -> tpe.ml_intern ist arg) () in - (GTacExt (tag, arg), GTypRef (Other tpe.ml_type, [])) + let arg, tpe = Flags.with_option Ltac_plugin.Tacintern.strict_check (fun () -> obj.ml_intern self ist arg) () in + (GTacExt (tag, arg), tpe) and intern_rec_with_constraint env e exp = let loc = loc_of_tacexpr e in -- cgit v1.2.3 From ba61b133772d76e6ff3f93da2ab136afd2f5a867 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 3 Sep 2017 18:23:17 +0200 Subject: Allowing ML objects to return mere tactic expressions. --- src/tac2core.ml | 12 ++++++------ src/tac2env.ml | 6 +++++- src/tac2env.mli | 6 +++++- src/tac2intern.ml | 6 +++++- 4 files changed, 21 insertions(+), 9 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index b8bec37d04..e1aa6eb48c 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -779,7 +779,7 @@ let gtypref kn = GTypRef (Other kn, []) let intern_constr self ist c = let (_, (c, _)) = Genintern.intern Stdarg.wit_constr ist c in - (c, gtypref t_constr) + (GlbVal c, gtypref t_constr) let interp_constr flags ist c = let open Pretyping in @@ -821,7 +821,7 @@ let () = let interp _ id = return (ValExt (Value.val_ident, id)) in let print _ id = str "ident:(" ++ Id.print id ++ str ")" in let obj = { - ml_intern = (fun _ _ id -> id, gtypref t_ident); + ml_intern = (fun _ _ id -> GlbVal id, gtypref t_ident); ml_interp = interp; ml_subst = (fun _ id -> id); ml_print = print; @@ -831,7 +831,7 @@ let () = let () = let intern self ist c = let _, pat = Constrintern.intern_constr_pattern ist.Genintern.genv ~as_type:false c in - pat, gtypref t_pattern + GlbVal pat, gtypref t_pattern in let print env pat = str "pattern:(" ++ Printer.pr_lconstr_pattern_env env Evd.empty pat ++ str ")" in let interp _ c = return (ValExt (Value.val_pattern, c)) in @@ -846,14 +846,14 @@ let () = let () = let intern self ist qid = match qid with | Libnames.Ident (_, id) -> - Globnames.VarRef id, gtypref t_reference + GlbVal (Globnames.VarRef id), gtypref t_reference | Libnames.Qualid (loc, qid) -> let gr = try Nametab.locate qid with Not_found -> Nametab.error_global_not_found ?loc qid in - gr, gtypref t_reference + GlbVal gr, gtypref t_reference in let subst s c = Globnames.subst_global_reference s c in let interp _ gr = return (Value.of_reference gr) in @@ -875,7 +875,7 @@ let () = let extra = Tac2intern.drop_ltac2_env ist.Genintern.extra in let ist = { ist with Genintern.extra } in let _, tac = Genintern.intern Ltac_plugin.Tacarg.wit_tactic ist tac in - tac, gtypref t_unit + GlbVal tac, gtypref t_unit in let interp _ tac = (** FUCK YOU API *) diff --git a/src/tac2env.ml b/src/tac2env.ml index 56fd55ee84..5a817df713 100644 --- a/src/tac2env.ml +++ b/src/tac2env.ml @@ -251,10 +251,14 @@ let shortest_qualid_of_projection kn = let sp = KNmap.find kn tab.tab_proj_rev in KnTab.shortest_qualid Id.Set.empty sp tab.tab_proj +type 'a or_glb_tacexpr = +| GlbVal of 'a +| GlbTacexpr of glb_tacexpr + type ('a, 'b, 'r) intern_fun = Genintern.glob_sign -> 'a -> 'b * 'r glb_typexpr type ('a, 'b) ml_object = { - ml_intern : 'r. (raw_tacexpr, glb_tacexpr, 'r) intern_fun -> ('a, 'b, 'r) intern_fun; + ml_intern : 'r. (raw_tacexpr, glb_tacexpr, 'r) intern_fun -> ('a, 'b or_glb_tacexpr, 'r) intern_fun; ml_subst : Mod_subst.substitution -> 'b -> 'b; ml_interp : environment -> 'b -> valexpr Proofview.tactic; ml_print : Environ.env -> 'b -> Pp.t; diff --git a/src/tac2env.mli b/src/tac2env.mli index 15664db756..eb18dc8e39 100644 --- a/src/tac2env.mli +++ b/src/tac2env.mli @@ -105,10 +105,14 @@ val interp_primitive : ml_tactic_name -> ml_tactic (** {5 ML primitive types} *) +type 'a or_glb_tacexpr = +| GlbVal of 'a +| GlbTacexpr of glb_tacexpr + type ('a, 'b, 'r) intern_fun = Genintern.glob_sign -> 'a -> 'b * 'r glb_typexpr type ('a, 'b) ml_object = { - ml_intern : 'r. (raw_tacexpr, glb_tacexpr, 'r) intern_fun -> ('a, 'b, 'r) intern_fun; + ml_intern : 'r. (raw_tacexpr, glb_tacexpr, 'r) intern_fun -> ('a, 'b or_glb_tacexpr, 'r) intern_fun; ml_subst : Mod_subst.substitution -> 'b -> 'b; ml_interp : environment -> 'b -> valexpr Proofview.tactic; ml_print : Environ.env -> 'b -> Pp.t; diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 5d2fc2b47b..2b234d7aec 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -789,7 +789,11 @@ let rec intern_rec env = function let ist = empty_glob_sign genv in let ist = { ist with extra = Store.set ist.extra ltac2_env env } in let arg, tpe = Flags.with_option Ltac_plugin.Tacintern.strict_check (fun () -> obj.ml_intern self ist arg) () in - (GTacExt (tag, arg), tpe) + let e = match arg with + | GlbVal arg -> GTacExt (tag, arg) + | GlbTacexpr e -> e + in + (e, tpe) and intern_rec_with_constraint env e exp = let loc = loc_of_tacexpr e in -- cgit v1.2.3 From 0b21a350f27d723a8f55a448be5ffde4841d21ad Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 3 Sep 2017 18:33:30 +0200 Subject: Uniform handling of locations in the various AST. --- src/g_ltac2.ml4 | 95 ++++++++++------- src/tac2core.ml | 14 ++- src/tac2entries.ml | 39 ++++--- src/tac2expr.mli | 46 ++++---- src/tac2intern.ml | 307 +++++++++++++++++++++++++---------------------------- src/tac2intern.mli | 8 +- src/tac2quote.ml | 76 ++++++------- 7 files changed, 288 insertions(+), 297 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index bfd2ab1a11..338711e79c 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -86,19 +86,19 @@ let tac2mode = Gram.entry_create "vernac:ltac2_command" (** FUCK YOU API *) let ltac1_expr = (Obj.magic Pltac.tactic_expr : Tacexpr.raw_tactic_expr Gram.entry) -let inj_wit wit loc x = CTacExt (loc, wit, x) +let inj_wit wit loc x = Loc.tag ~loc @@ CTacExt (wit, x) let inj_open_constr loc c = inj_wit Tac2env.wit_open_constr loc c let inj_pattern loc c = inj_wit Tac2env.wit_pattern loc c let inj_reference loc c = inj_wit Tac2env.wit_reference loc c let inj_ltac1 loc e = inj_wit Tac2env.wit_ltac1 loc e -let pattern_of_qualid loc id = - if Tac2env.is_constructor (snd id) then CPatRef (loc, RelId id, []) +let pattern_of_qualid ?loc id = + if Tac2env.is_constructor (snd id) then Loc.tag ?loc @@ CPatRef (RelId id, []) else let (dp, id) = Libnames.repr_qualid (snd id) in - if DirPath.is_empty dp then CPatVar (Some loc, Name id) + if DirPath.is_empty dp then Loc.tag ?loc @@ CPatVar (Name id) else - CErrors.user_err ~loc (Pp.str "Syntax error") + CErrors.user_err ?loc (Pp.str "Syntax error") GEXTEND Gram GLOBAL: tac2expr tac2type tac2def_val tac2def_typ tac2def_ext tac2def_syn @@ -107,53 +107,62 @@ GEXTEND Gram [ "1" LEFTA [ id = Prim.qualid; pl = LIST1 tac2pat LEVEL "0" -> if Tac2env.is_constructor (snd id) then - CPatRef (!@loc, RelId id, pl) + Loc.tag ~loc:!@loc @@ CPatRef (RelId id, pl) else CErrors.user_err ~loc:!@loc (Pp.str "Syntax error") - | id = Prim.qualid -> pattern_of_qualid !@loc id - | "["; "]" -> CPatRef (!@loc, AbsKn (Other Tac2core.Core.c_nil), []) + | id = Prim.qualid -> pattern_of_qualid ~loc:!@loc id + | "["; "]" -> Loc.tag ~loc:!@loc @@ CPatRef (AbsKn (Other Tac2core.Core.c_nil), []) | p1 = tac2pat; "::"; p2 = tac2pat -> - CPatRef (!@loc, AbsKn (Other Tac2core.Core.c_cons), [p1; p2]) + Loc.tag ~loc:!@loc @@ CPatRef (AbsKn (Other Tac2core.Core.c_cons), [p1; p2]) ] | "0" - [ "_" -> CPatVar (Some !@loc, Anonymous) - | "()" -> CPatRef (!@loc, AbsKn (Tuple 0), []) - | id = Prim.qualid -> pattern_of_qualid !@loc id + [ "_" -> Loc.tag ~loc:!@loc @@ CPatVar Anonymous + | "()" -> Loc.tag ~loc:!@loc @@ CPatRef (AbsKn (Tuple 0), []) + | id = Prim.qualid -> pattern_of_qualid ~loc:!@loc id | "("; pl = LIST0 tac2pat LEVEL "1" SEP ","; ")" -> - CPatRef (!@loc, AbsKn (Tuple (List.length pl)), pl) ] + Loc.tag ~loc:!@loc @@ CPatRef (AbsKn (Tuple (List.length pl)), pl) ] ] ; tac2expr: [ "6" RIGHTA - [ e1 = SELF; ";"; e2 = SELF -> CTacSeq (!@loc, e1, e2) ] + [ e1 = SELF; ";"; e2 = SELF -> Loc.tag ~loc:!@loc @@ CTacSeq (e1, e2) ] | "5" - [ "fun"; it = LIST1 input_fun ; "=>"; body = tac2expr LEVEL "6" -> CTacFun (!@loc, it, body) + [ "fun"; it = LIST1 input_fun ; "=>"; body = tac2expr LEVEL "6" -> + Loc.tag ~loc:!@loc @@ CTacFun (it, body) | "let"; isrec = rec_flag; lc = LIST1 let_clause SEP "with"; "in"; - e = tac2expr LEVEL "6" -> CTacLet (!@loc, isrec, lc, e) + e = tac2expr LEVEL "6" -> + Loc.tag ~loc:!@loc @@ CTacLet (isrec, lc, e) | "match"; e = tac2expr LEVEL "5"; "with"; bl = branches; "end" -> - CTacCse (!@loc, e, bl) + Loc.tag ~loc:!@loc @@ CTacCse (e, bl) ] | "4" LEFTA [ ] | "::" RIGHTA [ e1 = tac2expr; "::"; e2 = tac2expr -> - CTacApp (!@loc, CTacCst (!@loc, AbsKn (Other Tac2core.Core.c_cons)), [e1; e2]) + Loc.tag ~loc:!@loc @@ CTacApp (Loc.tag ~loc:!@loc @@ CTacCst (AbsKn (Other Tac2core.Core.c_cons)), [e1; e2]) ] | [ e0 = SELF; ","; el = LIST1 NEXT SEP "," -> let el = e0 :: el in - CTacApp (!@loc, CTacCst (!@loc, AbsKn (Tuple (List.length el))), el) ] + Loc.tag ~loc:!@loc @@ CTacApp (Loc.tag ~loc:!@loc @@ CTacCst (AbsKn (Tuple (List.length el))), el) ] | "1" LEFTA - [ e = tac2expr; el = LIST1 tac2expr LEVEL "0" -> CTacApp (!@loc, e, el) - | e = SELF; ".("; qid = Prim.qualid; ")" -> CTacPrj (!@loc, e, RelId qid) - | e = SELF; ".("; qid = Prim.qualid; ")"; ":="; r = tac2expr LEVEL "5" -> CTacSet (!@loc, e, RelId qid, r) ] + [ e = tac2expr; el = LIST1 tac2expr LEVEL "0" -> + Loc.tag ~loc:!@loc @@ CTacApp (e, el) + | e = SELF; ".("; qid = Prim.qualid; ")" -> + Loc.tag ~loc:!@loc @@ CTacPrj (e, RelId qid) + | e = SELF; ".("; qid = Prim.qualid; ")"; ":="; r = tac2expr LEVEL "5" -> + Loc.tag ~loc:!@loc @@ CTacSet (e, RelId qid, r) ] | "0" [ "("; a = SELF; ")" -> a - | "("; a = SELF; ":"; t = tac2type; ")" -> CTacCnv (!@loc, a, t) - | "()" -> CTacCst (!@loc, AbsKn (Tuple 0)) - | "("; ")" -> CTacCst (!@loc, AbsKn (Tuple 0)) + | "("; a = SELF; ":"; t = tac2type; ")" -> + Loc.tag ~loc:!@loc @@ CTacCnv (a, t) + | "()" -> + Loc.tag ~loc:!@loc @@ CTacCst (AbsKn (Tuple 0)) + | "("; ")" -> + Loc.tag ~loc:!@loc @@ CTacCst (AbsKn (Tuple 0)) | "["; a = LIST0 tac2expr LEVEL "5" SEP ";"; "]" -> Tac2quote.of_list ~loc:!@loc (fun x -> x) a - | "{"; a = tac2rec_fieldexprs; "}" -> CTacRec (!@loc, a) + | "{"; a = tac2rec_fieldexprs; "}" -> + Loc.tag ~loc:!@loc @@ CTacRec a | a = tactic_atom -> a ] ] ; @@ -178,10 +187,13 @@ GEXTEND Gram [ [ "'"; id = Prim.ident -> id ] ] ; tactic_atom: - [ [ n = Prim.integer -> CTacAtm (Loc.tag ~loc:!@loc (AtmInt n)) - | s = Prim.string -> CTacAtm (Loc.tag ~loc:!@loc (AtmStr s)) + [ [ n = Prim.integer -> Loc.tag ~loc:!@loc @@ CTacAtm (AtmInt n) + | s = Prim.string -> Loc.tag ~loc:!@loc @@ CTacAtm (AtmStr s) | id = Prim.qualid -> - if Tac2env.is_constructor (snd id) then CTacCst (!@loc, RelId id) else CTacRef (RelId id) + if Tac2env.is_constructor (snd id) then + Loc.tag ~loc:!@loc @@ CTacCst (RelId id) + else + Loc.tag ~loc:!@loc @@ CTacRef (RelId id) | "@"; id = Prim.ident -> Tac2quote.of_ident (Loc.tag ~loc:!@loc id) | "&"; id = lident -> Tac2quote.of_hyp ~loc:!@loc id | "'"; c = Constr.constr -> inj_open_constr !@loc c @@ -196,35 +208,38 @@ GEXTEND Gram let_clause: [ [ binder = let_binder; ":="; te = tac2expr -> let (pat, fn) = binder in - let te = match fn with None -> te | Some args -> CTacFun (!@loc, args, te) in + let te = match fn with + | None -> te + | Some args -> Loc.tag ~loc:!@loc @@ CTacFun (args, te) + in (pat, None, te) ] ] ; let_binder: [ [ pats = LIST1 input_fun -> match pats with - | [CPatVar _ as pat, None] -> (pat, None) - | (CPatVar (_, Name id) as pat, None) :: args -> (pat, Some args) + | [(_, CPatVar _) as pat, None] -> (pat, None) + | ((_, CPatVar (Name id)) as pat, None) :: args -> (pat, Some args) | [pat, None] -> (pat, None) | _ -> CErrors.user_err ~loc:!@loc (str "Invalid pattern") ] ] ; tac2type: [ "5" RIGHTA - [ t1 = tac2type; "->"; t2 = tac2type -> CTypArrow (!@loc, t1, t2) ] + [ t1 = tac2type; "->"; t2 = tac2type -> Loc.tag ~loc:!@loc @@ CTypArrow (t1, t2) ] | "2" [ t = tac2type; "*"; tl = LIST1 tac2type LEVEL "1" SEP "*" -> let tl = t :: tl in - CTypRef (!@loc, AbsKn (Tuple (List.length tl)), tl) ] + Loc.tag ~loc:!@loc @@ CTypRef (AbsKn (Tuple (List.length tl)), tl) ] | "1" LEFTA - [ t = SELF; qid = Prim.qualid -> CTypRef (!@loc, RelId qid, [t]) ] + [ t = SELF; qid = Prim.qualid -> Loc.tag ~loc:!@loc @@ CTypRef (RelId qid, [t]) ] | "0" [ "("; t = tac2type LEVEL "5"; ")" -> t - | id = typ_param -> CTypVar (Loc.tag ~loc:!@loc (Name id)) - | "_" -> CTypVar (Loc.tag ~loc:!@loc Anonymous) - | qid = Prim.qualid -> CTypRef (!@loc, RelId qid, []) + | id = typ_param -> Loc.tag ~loc:!@loc @@ CTypVar (Name id) + | "_" -> Loc.tag ~loc:!@loc @@ CTypVar Anonymous + | qid = Prim.qualid -> Loc.tag ~loc:!@loc @@ CTypRef (RelId qid, []) | "("; p = LIST1 tac2type LEVEL "5" SEP ","; ")"; qid = Prim.qualid -> - CTypRef (!@loc, RelId qid, p) ] + Loc.tag ~loc:!@loc @@ CTypRef (RelId qid, p) ] ]; locident: [ [ id = Prim.ident -> Loc.tag ~loc:!@loc id ] ] @@ -239,7 +254,7 @@ GEXTEND Gram ; tac2def_body: [ [ name = binder; it = LIST0 input_fun; ":="; e = tac2expr -> - let e = if List.is_empty it then e else CTacFun (!@loc, it, e) in + let e = if List.is_empty it then e else Loc.tag ~loc:!@loc @@ CTacFun (it, e) in (name, e) ] ] ; diff --git a/src/tac2core.ml b/src/tac2core.ml index e1aa6eb48c..db8f989768 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -938,20 +938,18 @@ let add_scope s f = let scope_fail () = CErrors.user_err (str "Invalid parsing token") -let dummy_loc = Loc.make_loc (-1, -1) - -let q_unit = CTacCst (dummy_loc, AbsKn (Tuple 0)) +let q_unit = Loc.tag @@ CTacCst (AbsKn (Tuple 0)) let rthunk e = let loc = Tac2intern.loc_of_tacexpr e in - let var = [CPatVar (Some loc, Anonymous), Some (CTypRef (loc, AbsKn (Other Core.t_unit), []))] in - CTacFun (loc, var, e) + let var = [Loc.tag ?loc @@ CPatVar Anonymous, Some (Loc.tag ?loc @@ CTypRef (AbsKn (Other Core.t_unit), []))] in + Loc.tag ?loc @@ CTacFun (var, e) let add_generic_scope s entry arg = let parse = function | [] -> let scope = Extend.Aentry entry in - let act x = CTacExt (dummy_loc, arg, x) in + let act x = Loc.tag @@ CTacExt (arg, x) in Tac2entries.ScopeRule (scope, act) | _ -> scope_fail () in @@ -1007,9 +1005,9 @@ let () = add_scope "opt" begin function let scope = Extend.Aopt scope in let act opt = match opt with | None -> - CTacCst (dummy_loc, AbsKn (Other Core.c_none)) + Loc.tag @@ CTacCst (AbsKn (Other Core.c_none)) | Some x -> - CTacApp (dummy_loc, CTacCst (dummy_loc, AbsKn (Other Core.c_some)), [act x]) + Loc.tag @@ CTacApp (Loc.tag @@ CTacCst (AbsKn (Other Core.c_some)), [act x]) in Tac2entries.ScopeRule (scope, act) | _ -> scope_fail () diff --git a/src/tac2entries.ml b/src/tac2entries.ml index a503a0ab4c..9c5d9a659b 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -263,8 +263,6 @@ let inTypExt : typext -> obj = (** Toplevel entries *) -let dummy_loc = Loc.make_loc (-1, -1) - let fresh_var avoid x = let bad id = Id.Set.mem id avoid || @@ -275,8 +273,8 @@ let fresh_var avoid x = (** Mangle recursive tactics *) let inline_rec_tactic tactics = let avoid = List.fold_left (fun accu ((_, id), _) -> Id.Set.add id accu) Id.Set.empty tactics in - let map (id, e) = match e with - | CTacFun (loc, pat, _) -> (id, pat, e) + let map (id, e) = match snd e with + | CTacFun (pat, _) -> (id, pat, e) | _ -> let loc, _ = id in user_err ?loc (str "Recursive tactic definitions must be functions") @@ -286,24 +284,24 @@ let inline_rec_tactic tactics = let fold_var (avoid, ans) (pat, _) = let id = fresh_var avoid "x" in let loc = loc_of_patexpr pat in - (Id.Set.add id avoid, Loc.tag ~loc id :: ans) + (Id.Set.add id avoid, Loc.tag ?loc id :: ans) in (** Fresh variables to abstract over the function patterns *) let _, vars = List.fold_left fold_var (avoid, []) pat in - let map_body ((loc, id), _, e) = CPatVar (loc, Name id), None, e in + let map_body ((loc, id), _, e) = (Loc.tag ?loc @@ CPatVar (Name id)), None, e in let bnd = List.map map_body tactics in let pat_of_id (loc, id) = - (CPatVar (loc, Name id), None) + ((Loc.tag ?loc @@ CPatVar (Name id)), None) in let var_of_id (loc, id) = let qid = (loc, qualid_of_ident id) in - CTacRef (RelId qid) + Loc.tag ?loc @@ CTacRef (RelId qid) in let loc0 = loc_of_tacexpr e in let vpat = List.map pat_of_id vars in let varg = List.map var_of_id vars in - let e = CTacLet (loc0, true, bnd, CTacApp (loc0, var_of_id id, varg)) in - (id, CTacFun (loc0, vpat, e)) + let e = Loc.tag ?loc:loc0 @@ CTacLet (true, bnd, Loc.tag ?loc:loc0 @@ CTacApp (var_of_id id, varg)) in + (id, Loc.tag ?loc:loc0 @@ CTacFun (vpat, e)) in List.map map tactics @@ -459,9 +457,8 @@ let register_open ?(local = false) (loc, qid) (params, def) = user_err ?loc (str "Type " ++ pr_qualid qid ++ str " is not an open type") in let () = - let loc = Option.default dummy_loc loc in if not (Int.equal (List.length params) tparams) then - Tac2intern.error_nparams_mismatch loc (List.length params) tparams + Tac2intern.error_nparams_mismatch ?loc (List.length params) tparams in match def with | CTydOpn -> () @@ -524,9 +521,9 @@ module ParseToken = struct let loc_of_token = function -| SexprStr (loc, _) -> Option.default dummy_loc loc -| SexprInt (loc, _) -> Option.default dummy_loc loc -| SexprRec (loc, _, _) -> loc +| SexprStr (loc, _) -> loc +| SexprInt (loc, _) -> loc +| SexprRec (loc, _, _) -> Some loc let parse_scope = function | SexprRec (_, (loc, Some id), toks) -> @@ -535,11 +532,11 @@ let parse_scope = function else CErrors.user_err ?loc (str "Unknown scope" ++ spc () ++ Nameops.pr_id id) | SexprStr (_, str) -> - let v_unit = CTacCst (dummy_loc, AbsKn (Tuple 0)) in + let v_unit = Loc.tag @@ CTacCst (AbsKn (Tuple 0)) in ScopeRule (Extend.Atoken (Tok.IDENT str), (fun _ -> v_unit)) | tok -> let loc = loc_of_token tok in - CErrors.user_err ~loc (str "Invalid parsing token") + CErrors.user_err ?loc (str "Invalid parsing token") let parse_token = function | SexprStr (_, s) -> TacTerm s @@ -549,7 +546,7 @@ let parse_token = function TacNonTerm (na, scope) | tok -> let loc = loc_of_token tok in - CErrors.user_err ~loc (str "Invalid parsing token") + CErrors.user_err ?loc (str "Invalid parsing token") end @@ -586,10 +583,10 @@ let perform_notation syn st = let mk loc args = let map (na, e) = let loc = loc_of_tacexpr e in - (CPatVar (Loc.tag ~loc na), None, e) + ((Loc.tag ?loc @@ CPatVar na), None, e) in let bnd = List.map map args in - CTacLet (loc, false, bnd, syn.synext_exp) + Loc.tag ~loc @@ CTacLet (false, bnd, syn.synext_exp) in let rule = Extend.Rule (rule, act mk) in let lev = match syn.synext_lev with @@ -793,7 +790,7 @@ let solve default tac = let call ~default e = let loc = loc_of_tacexpr e in let (e, t) = intern e in - let () = check_unit ~loc t in + let () = check_unit ?loc t in let tac = Tac2interp.interp Id.Map.empty e in solve default (Proofview.tclIGNORE tac) diff --git a/src/tac2expr.mli b/src/tac2expr.mli index ccff8e7756..1045063cd2 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -44,10 +44,12 @@ type 'a or_tuple = (** {5 Type syntax} *) -type raw_typexpr = -| CTypVar of Name.t located -| CTypArrow of Loc.t * raw_typexpr * raw_typexpr -| CTypRef of Loc.t * type_constant or_tuple or_relid * raw_typexpr list +type raw_typexpr_r = +| CTypVar of Name.t +| CTypArrow of raw_typexpr * raw_typexpr +| CTypRef of type_constant or_tuple or_relid * raw_typexpr list + +and raw_typexpr = raw_typexpr_r located type raw_typedef = | CTydDef of raw_typexpr option @@ -87,24 +89,28 @@ type atom = | AtmStr of string (** Tactic expressions *) -type raw_patexpr = -| CPatVar of Name.t located -| CPatRef of Loc.t * ltac_constructor or_tuple or_relid * raw_patexpr list +type raw_patexpr_r = +| CPatVar of Name.t +| CPatRef of ltac_constructor or_tuple or_relid * raw_patexpr list + +and raw_patexpr = raw_patexpr_r located -type raw_tacexpr = -| CTacAtm of atom located +type raw_tacexpr_r = +| CTacAtm of atom | CTacRef of tacref or_relid -| CTacCst of Loc.t * ltac_constructor or_tuple or_relid -| CTacFun of Loc.t * (raw_patexpr * raw_typexpr option) list * raw_tacexpr -| CTacApp of Loc.t * raw_tacexpr * raw_tacexpr list -| CTacLet of Loc.t * rec_flag * (raw_patexpr * raw_typexpr option * raw_tacexpr) list * raw_tacexpr -| CTacCnv of Loc.t * raw_tacexpr * raw_typexpr -| CTacSeq of Loc.t * raw_tacexpr * raw_tacexpr -| CTacCse of Loc.t * raw_tacexpr * raw_taccase list -| CTacRec of Loc.t * raw_recexpr -| CTacPrj of Loc.t * raw_tacexpr * ltac_projection or_relid -| CTacSet of Loc.t * raw_tacexpr * ltac_projection or_relid * raw_tacexpr -| CTacExt : Loc.t * ('a, _) Tac2dyn.Arg.tag * 'a -> raw_tacexpr +| CTacCst of ltac_constructor or_tuple or_relid +| CTacFun of (raw_patexpr * raw_typexpr option) list * raw_tacexpr +| CTacApp of raw_tacexpr * raw_tacexpr list +| CTacLet of rec_flag * (raw_patexpr * raw_typexpr option * raw_tacexpr) list * raw_tacexpr +| CTacCnv of raw_tacexpr * raw_typexpr +| CTacSeq of raw_tacexpr * raw_tacexpr +| CTacCse of raw_tacexpr * raw_taccase list +| CTacRec of raw_recexpr +| CTacPrj of raw_tacexpr * ltac_projection or_relid +| CTacSet of raw_tacexpr * ltac_projection or_relid * raw_tacexpr +| CTacExt : ('a, _) Tac2dyn.Arg.tag * 'a -> raw_tacexpr_r + +and raw_tacexpr = raw_tacexpr_r located and raw_taccase = raw_patexpr * raw_tacexpr diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 2b234d7aec..c1fd281808 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -187,36 +187,18 @@ let push_name id t env = match id with | Anonymous -> env | Name id -> { env with env_var = Id.Map.add id t env.env_var } -let dummy_loc = Loc.make_loc (-1, -1) - -let loc_of_tacexpr = function -| CTacAtm (loc, _) -> Option.default dummy_loc loc -| CTacRef (RelId (loc, _)) -> Option.default dummy_loc loc -| CTacRef (AbsKn _) -> dummy_loc -| CTacCst (loc, _) -> loc -| CTacFun (loc, _, _) -> loc -| CTacApp (loc, _, _) -> loc -| CTacLet (loc, _, _, _) -> loc -| CTacCnv (loc, _, _) -> loc -| CTacSeq (loc, _, _) -> loc -| CTacCse (loc, _, _) -> loc -| CTacRec (loc, _) -> loc -| CTacPrj (loc, _, _) -> loc -| CTacSet (loc, _, _, _) -> loc -| CTacExt (loc, _, _) -> loc - -let loc_of_patexpr = function -| CPatVar (loc, _) -> Option.default dummy_loc loc -| CPatRef (loc, _, _) -> loc - -let error_nargs_mismatch loc kn nargs nfound = +let loc_of_tacexpr (loc, _) : Loc.t option = loc + +let loc_of_patexpr (loc, _) : Loc.t option = loc + +let error_nargs_mismatch ?loc kn nargs nfound = let cstr = Tac2env.shortest_qualid_of_constructor kn in - user_err ~loc (str "Constructor " ++ pr_qualid cstr ++ str " expects " ++ + user_err ?loc (str "Constructor " ++ pr_qualid cstr ++ str " expects " ++ int nargs ++ str " arguments, but is applied to " ++ int nfound ++ str " arguments") -let error_nparams_mismatch loc nargs nfound = - user_err ~loc (str "Type expects " ++ int nargs ++ +let error_nparams_mismatch ?loc nargs nfound = + user_err ?loc (str "Type expects " ++ int nargs ++ str " arguments, but is applied to " ++ int nfound ++ str " arguments") @@ -226,10 +208,10 @@ let rec subst_type subst (t : 'a glb_typexpr) = match t with | GTypRef (qid, args) -> GTypRef (qid, List.map (fun t -> subst_type subst t) args) -let rec intern_type env (t : raw_typexpr) : UF.elt glb_typexpr = match t with -| CTypVar (loc, Name id) -> GTypVar (get_alias (Loc.tag ?loc id) env) -| CTypVar (_, Anonymous) -> GTypVar (fresh_id env) -| CTypRef (loc, rel, args) -> +let rec intern_type env ((loc, t) : raw_typexpr) : UF.elt glb_typexpr = match t with +| CTypVar (Name id) -> GTypVar (get_alias (Loc.tag ?loc id) env) +| CTypVar Anonymous -> GTypVar (fresh_id env) +| CTypRef (rel, args) -> let (kn, nparams) = match rel with | RelId (loc, qid) -> let (dp, id) = repr_qualid qid in @@ -255,7 +237,7 @@ let rec intern_type env (t : raw_typexpr) : UF.elt glb_typexpr = match t with if not (Int.equal nparams nargs) then let loc, qid = match rel with | RelId lid -> lid - | AbsKn (Other kn) -> Some loc, shortest_qualid_of_type kn + | AbsKn (Other kn) -> loc, shortest_qualid_of_type kn | AbsKn (Tuple _) -> assert false in user_err ?loc (strbrk "The type constructor " ++ pr_qualid qid ++ @@ -263,7 +245,7 @@ let rec intern_type env (t : raw_typexpr) : UF.elt glb_typexpr = match t with applied to " ++ int nargs ++ strbrk "argument(s)") in GTypRef (kn, List.map (fun t -> intern_type env t) args) -| CTypArrow (loc, t1, t2) -> GTypArrow (intern_type env t1, intern_type env t2) +| CTypArrow (t1, t2) -> GTypArrow (intern_type env t1, intern_type env t2) let fresh_type_scheme env (t : type_scheme) : UF.elt glb_typexpr = let (n, t) = t in @@ -387,7 +369,7 @@ let unify_arrow ?loc env ft args = let rec iter ft args is_fun = match kind env ft, args with | t, [] -> t | GTypArrow (t1, ft), (loc, t2) :: args -> - let () = unify ~loc env t2 t1 in + let () = unify ?loc env t2 t1 in iter ft args true | GTypVar id, (_, t) :: args -> let ft = GTypVar (fresh_id env) in @@ -492,19 +474,19 @@ let check_elt_unit loc env t = | GTypRef (Tuple 0, []) -> true | GTypRef _ -> false in - if not maybe_unit then warn_not_unit ~loc () + if not maybe_unit then warn_not_unit ?loc () let check_elt_empty loc env t = match kind env t with | GTypVar _ -> - user_err ~loc (str "Cannot infer an empty type for this expression") + user_err ?loc (str "Cannot infer an empty type for this expression") | GTypArrow _ | GTypRef (Tuple _, _) -> - user_err ~loc (str "Type " ++ pr_glbtype env t ++ str " is not an empty type") + user_err ?loc (str "Type " ++ pr_glbtype env t ++ str " is not an empty type") | GTypRef (Other kn, _) -> let def = Tac2env.interp_type kn in match def with | _, GTydAlg { galg_constructors = [] } -> kn | _ -> - user_err ~loc (str "Type " ++ pr_glbtype env t ++ str " is not an empty type") + user_err ?loc (str "Type " ++ pr_glbtype env t ++ str " is not an empty type") let check_unit ?loc t = let env = empty_env () in @@ -520,7 +502,7 @@ let check_unit ?loc t = let check_redundant_clause = function | [] -> () -| (p, _) :: _ -> warn_redundant_clause ~loc:(loc_of_patexpr p) () +| (p, _) :: _ -> warn_redundant_clause ?loc:(loc_of_patexpr p) () let get_variable0 mem var = match var with | RelId (loc, qid) -> @@ -576,9 +558,9 @@ type glb_patexpr = | GPatVar of Name.t | GPatRef of ltac_constructor or_tuple * glb_patexpr list -let rec intern_patexpr env = function -| CPatVar (_, na) -> GPatVar na -| CPatRef (_, qid, pl) -> +let rec intern_patexpr env (_, pat) = match pat with +| CPatVar na -> GPatVar na +| CPatRef (qid, pl) -> let kn = get_constructor env qid in GPatRef (kn, List.map (fun p -> intern_patexpr env p) pl) @@ -619,27 +601,27 @@ let add_name accu = function | Name id -> Id.Set.add id accu | Anonymous -> accu -let rec ids_of_pattern accu = function -| CPatVar (_, Anonymous) -> accu -| CPatVar (_, Name id) -> Id.Set.add id accu -| CPatRef (_, _, pl) -> +let rec ids_of_pattern accu (_, pat) = match pat with +| CPatVar Anonymous -> accu +| CPatVar (Name id) -> Id.Set.add id accu +| CPatRef (_, pl) -> List.fold_left ids_of_pattern accu pl let loc_of_relid = function -| RelId (loc, _) -> Option.default dummy_loc loc -| AbsKn _ -> dummy_loc +| RelId (loc, _) -> loc +| AbsKn _ -> None (** Expand pattern: [p => t] becomes [x => match x with p => t end] *) let expand_pattern avoid bnd = let fold (avoid, bnd) (pat, t) = - let na, expand = match pat with - | CPatVar (_, na) -> + let na, expand = match snd pat with + | CPatVar na -> (** Don't expand variable patterns *) na, None | _ -> let loc = loc_of_patexpr pat in let id = fresh_var avoid in - let qid = RelId (Loc.tag ~loc (qualid_of_ident id)) in + let qid = RelId (Loc.tag ?loc (qualid_of_ident id)) in Name id, Some qid in let avoid = ids_of_pattern avoid pat in @@ -649,7 +631,9 @@ let expand_pattern avoid bnd = let (_, bnd) = List.fold_left fold (avoid, []) bnd in let fold e (na, pat, expand) = match expand with | None -> e - | Some qid -> CTacCse (loc_of_relid qid, CTacRef qid, [pat, e]) + | Some qid -> + let loc = loc_of_relid qid in + Loc.tag ?loc @@ CTacCse (Loc.tag ?loc @@ CTacRef qid, [pat, e]) in let expand e = List.fold_left fold e bnd in let nas = List.rev_map (fun (na, _, _) -> na) bnd in @@ -659,8 +643,8 @@ let is_alias env qid = match get_variable env qid with | ArgArg (TacAlias _) -> true | ArgVar _ | (ArgArg (TacConstant _)) -> false -let rec intern_rec env = function -| CTacAtm (_, atm) -> intern_atm env atm +let rec intern_rec env (loc, e) = match e with +| CTacAtm atm -> intern_atm env atm | CTacRef qid -> begin match get_variable env qid with | ArgVar (_, id) -> @@ -673,10 +657,10 @@ let rec intern_rec env = function let e = Tac2env.interp_alias kn in intern_rec env e end -| CTacCst (loc, qid) -> +| CTacCst qid -> let kn = get_constructor env qid in intern_constructor env loc kn [] -| CTacFun (loc, bnd, e) -> +| CTacFun (bnd, e) -> let map (_, t) = match t with | None -> GTypVar (fresh_id env) | Some t -> intern_type env t @@ -687,10 +671,10 @@ let rec intern_rec env = function let (e, t) = intern_rec env (exp e) in let t = List.fold_right (fun t accu -> GTypArrow (t, accu)) tl t in (GTacFun (nas, e), t) -| CTacApp (loc, CTacCst (_, qid), args) -> +| CTacApp ((loc, CTacCst qid), args) -> let kn = get_constructor env qid in intern_constructor env loc kn args -| CTacApp (loc, CTacRef qid, args) when is_alias env qid -> +| CTacApp ((_, CTacRef qid), args) when is_alias env qid -> let kn = match get_variable env qid with | ArgArg (TacAlias kn) -> kn | ArgVar _ | (ArgArg (TacConstant _)) -> assert false @@ -699,12 +683,12 @@ let rec intern_rec env = function let map arg = (** Thunk alias arguments *) let loc = loc_of_tacexpr arg in - let var = [CPatVar (Some loc, Anonymous), Some (CTypRef (loc, AbsKn (Tuple 0), []))] in - CTacFun (loc, var, arg) + let var = [Loc.tag ?loc @@ CPatVar Anonymous, Some (Loc.tag ?loc @@ CTypRef (AbsKn (Tuple 0), []))] in + Loc.tag ?loc @@ CTacFun (var, arg) in let args = List.map map args in - intern_rec env (CTacApp (loc, e, args)) -| CTacApp (loc, f, args) -> + intern_rec env (Loc.tag ?loc @@ CTacApp (e, args)) +| CTacApp (f, args) -> let loc = loc_of_tacexpr f in let (f, ft) = intern_rec env f in let fold arg (args, t) = @@ -713,9 +697,9 @@ let rec intern_rec env = function (arg :: args, (loc, argt) :: t) in let (args, t) = List.fold_right fold args ([], []) in - let ret = unify_arrow ~loc env ft t in + let ret = unify_arrow ?loc env ft t in (GTacApp (f, args), ret) -| CTacLet (loc, is_rec, el, e) -> +| CTacLet (is_rec, el, e) -> let fold accu (pat, _, e) = let ids = ids_of_pattern Id.Set.empty pat in let common = Id.Set.inter ids accu in @@ -723,39 +707,39 @@ let rec intern_rec env = function else let id = Id.Set.choose common in let loc = loc_of_patexpr pat in - user_err ~loc (str "Variable " ++ Id.print id ++ str " is bound several \ + user_err ?loc (str "Variable " ++ Id.print id ++ str " is bound several \ times in this matching") in let ids = List.fold_left fold Id.Set.empty el in if is_rec then intern_let_rec env loc ids el e else intern_let env loc ids el e -| CTacCnv (loc, e, tc) -> +| CTacCnv (e, tc) -> let (e, t) = intern_rec env e in let tc = intern_type env tc in - let () = unify ~loc env t tc in + let () = unify ?loc env t tc in (e, tc) -| CTacSeq (loc, e1, e2) -> +| CTacSeq (e1, e2) -> let loc1 = loc_of_tacexpr e1 in let (e1, t1) = intern_rec env e1 in let (e2, t2) = intern_rec env e2 in let () = check_elt_unit loc1 env t1 in (GTacLet (false, [Anonymous, e1], e2), t2) -| CTacCse (loc, e, pl) -> +| CTacCse (e, pl) -> intern_case env loc e pl -| CTacRec (loc, fs) -> +| CTacRec fs -> intern_record env loc fs -| CTacPrj (loc, e, proj) -> +| CTacPrj (e, proj) -> let pinfo = get_projection proj in let loc = loc_of_tacexpr e in let (e, t) = intern_rec env e in let subst = Array.init pinfo.pdata_prms (fun _ -> fresh_id env) in let params = Array.map_to_list (fun i -> GTypVar i) subst in let exp = GTypRef (Other pinfo.pdata_type, params) in - let () = unify ~loc env t exp in + let () = unify ?loc env t exp in let substf i = GTypVar subst.(i) in let ret = subst_type substf pinfo.pdata_ptyp in (GTacPrj (pinfo.pdata_type, e, pinfo.pdata_indx), ret) -| CTacSet (loc, e, proj, r) -> +| CTacSet (e, proj, r) -> let pinfo = get_projection proj in let () = if not pinfo.pdata_mutb then @@ -773,7 +757,7 @@ let rec intern_rec env = function let ret = subst_type substf pinfo.pdata_ptyp in let r = intern_rec_with_constraint env r ret in (GTacSet (pinfo.pdata_type, e, pinfo.pdata_indx, r), GTypRef (Tuple 0, [])) -| CTacExt (loc, tag, arg) -> +| CTacExt (tag, arg) -> let open Genintern in let self ist e = let env = match Store.get ist.extra ltac2_env with @@ -798,7 +782,7 @@ let rec intern_rec env = function and intern_rec_with_constraint env e exp = let loc = loc_of_tacexpr e in let (e, t) = intern_rec env e in - let () = unify ~loc env t exp in + let () = unify ?loc env t exp in e and intern_let env loc ids el e = @@ -827,11 +811,11 @@ and intern_let env loc ids el e = and intern_let_rec env loc ids el e = let map env (pat, t, e) = - let loc, na = match pat with + let (loc, pat) = pat in + let na = match pat with | CPatVar na -> na | CPatRef _ -> - let loc = loc_of_patexpr pat in - user_err ~loc (str "This kind of pattern is forbidden in let-rec bindings") + user_err ?loc (str "This kind of pattern is forbidden in let-rec bindings") in let id = fresh_id env in let env = push_name na (monomorphic (GTypVar id)) env in @@ -843,7 +827,7 @@ and intern_let_rec env loc ids el e = let (e, t) = intern_rec env e in let () = if not (is_rec_rhs e) then - user_err ~loc:loc_e (str "This kind of expression is not allowed as \ + user_err ?loc:loc_e (str "This kind of expression is not allowed as \ right-hand side of a recursive binding") in let () = unify ?loc env t (GTypVar id) in @@ -881,7 +865,7 @@ and intern_case env loc e pl = (GTacCse (e', Other kn, [||], [||]), GTypVar r) | PKind_variant kn -> let subst, tc = fresh_reftype env kn in - let () = unify ~loc:(loc_of_tacexpr e) env t tc in + let () = unify ?loc:(loc_of_tacexpr e) env t tc in let (nconst, nnonconst, arities) = match kn with | Tuple 0 -> 1, 0, [0] | Tuple n -> 0, 1, [n] @@ -897,9 +881,11 @@ and intern_case env loc e pl = let rec intern_branch = function | [] -> () | (pat, br) :: rem -> - let tbr = match pat with - | CPatVar (loc, Name _) -> todo ?loc () - | CPatVar (_, Anonymous) -> + let tbr = match snd pat with + | CPatVar (Name _) -> + let loc = loc_of_patexpr pat in + todo ?loc () + | CPatVar Anonymous -> let () = check_redundant_clause rem in let (br', brT) = intern_rec env br in (** Fill all remaining branches *) @@ -919,7 +905,8 @@ and intern_case env loc e pl = in let _ = List.fold_left fill (0, 0) arities in brT - | CPatRef (loc, qid, args) -> + | CPatRef (qid, args) -> + let loc = loc_of_patexpr pat in let knc = get_constructor env qid in let kn', index, arity = match knc with | Tuple n -> Tuple n, 0, List.init n (fun i -> GTypVar i) @@ -930,11 +917,11 @@ and intern_case env loc e pl = in let () = if not (eq_or_tuple KerName.equal kn kn') then - invalid_pattern ~loc kn kn' + invalid_pattern ?loc kn kn' in - let get_id = function - | CPatVar (_, na) -> na - | p -> todo ~loc:(loc_of_patexpr p) () + let get_id pat = match pat with + | _, CPatVar na -> na + | loc, _ -> todo ?loc () in let ids = List.map get_id args in let nids = List.length ids in @@ -942,7 +929,7 @@ and intern_case env loc e pl = let () = match knc with | Tuple n -> assert (n == nids) | Other knc -> - if not (Int.equal nids nargs) then error_nargs_mismatch loc knc nargs nids + if not (Int.equal nids nargs) then error_nargs_mismatch ?loc knc nargs nids in let fold env id tpe = (** Instantiate all arguments *) @@ -955,15 +942,15 @@ and intern_case env loc e pl = let () = if List.is_empty args then if Option.is_empty const.(index) then const.(index) <- Some br' - else warn_redundant_clause ~loc () + else warn_redundant_clause ?loc () else let ids = Array.of_list ids in if Option.is_empty nonconst.(index) then nonconst.(index) <- Some (ids, br') - else warn_redundant_clause ~loc () + else warn_redundant_clause ?loc () in brT in - let () = unify ~loc:(loc_of_tacexpr br) env tbr ret in + let () = unify ?loc:(loc_of_tacexpr br) env tbr ret in intern_branch rem in let () = intern_branch pl in @@ -971,7 +958,7 @@ and intern_case env loc e pl = | None -> let kn = match kn with Other kn -> kn | _ -> assert false in let cstr = pr_internal_constructor kn n is_const in - user_err ~loc (str "Unhandled match case for constructor " ++ cstr) + user_err ?loc (str "Unhandled match case for constructor " ++ cstr) | Some x -> x in let const = Array.mapi (fun i o -> map i true o) const in @@ -980,11 +967,11 @@ and intern_case env loc e pl = (ce, ret) | PKind_open kn -> let subst, tc = fresh_reftype env (Other kn) in - let () = unify ~loc:(loc_of_tacexpr e) env t tc in + let () = unify ?loc:(loc_of_tacexpr e) env t tc in let ret = GTypVar (fresh_id env) in let rec intern_branch map = function | [] -> - user_err ~loc (str "Missing default case") + user_err ?loc (str "Missing default case") | (pat, br) :: rem -> match intern_patexpr env pat with | GPatVar na -> @@ -997,23 +984,23 @@ and intern_case env loc e pl = let get = function | GPatVar na -> na | GPatRef _ -> - user_err ~loc (str "TODO: Unhandled match case") (** FIXME *) + user_err ?loc (str "TODO: Unhandled match case") (** FIXME *) in let loc = loc_of_patexpr pat in let knc = match knc with | Other knc -> knc - | Tuple n -> invalid_pattern ~loc (Other kn) (Tuple n) + | Tuple n -> invalid_pattern ?loc (Other kn) (Tuple n) in let ids = List.map get args in let data = Tac2env.interp_constructor knc in let () = if not (KerName.equal kn data.cdata_type) then - invalid_pattern ~loc (Other kn) (Other data.cdata_type) + invalid_pattern ?loc (Other kn) (Other data.cdata_type) in let nids = List.length ids in let nargs = List.length data.cdata_args in let () = - if not (Int.equal nids nargs) then error_nargs_mismatch loc knc nargs nids + if not (Int.equal nids nargs) then error_nargs_mismatch ?loc knc nargs nids in let fold env id tpe = (** Instantiate all arguments *) @@ -1025,7 +1012,7 @@ and intern_case env loc e pl = let br' = intern_rec_with_constraint nenv br ret in let map = if KNmap.mem knc map then - let () = warn_redundant_clause ~loc () in + let () = warn_redundant_clause ?loc () in map else KNmap.add knc (Anonymous, Array.of_list ids, br') map @@ -1053,7 +1040,7 @@ and intern_constructor env loc kn args = match kn with | None -> (GTacOpn (kn, args), ans) else - error_nargs_mismatch loc kn nargs (List.length args) + error_nargs_mismatch ?loc kn nargs (List.length args) | Tuple n -> assert (Int.equal n (List.length args)); let types = List.init n (fun i -> GTypVar (fresh_id env)) in @@ -1073,7 +1060,7 @@ and intern_record env loc fs = in let fs = List.map map fs in let kn = match fs with - | [] -> user_err ~loc (str "Cannot infer the corresponding record type") + | [] -> user_err ?loc (str "Cannot infer the corresponding record type") | (_, proj, _) :: _ -> proj.pdata_type in let params, typdef = match Tac2env.interp_type kn with @@ -1104,7 +1091,7 @@ and intern_record env loc fs = | None -> () | Some i -> let (field, _, _) = List.nth typdef i in - user_err ~loc (str "Field " ++ Id.print field ++ str " is undefined") + user_err ?loc (str "Field " ++ Id.print field ++ str " is undefined") in let args = Array.map_to_list Option.get args in let tparam = List.init params (fun i -> GTypVar subst.(i)) in @@ -1204,18 +1191,18 @@ let get_projection0 var = match var with kn | AbsKn kn -> kn -let rec globalize ids e = match e with +let rec globalize ids (loc, er as e) = match er with | CTacAtm _ -> e | CTacRef ref -> let mem id = Id.Set.mem id ids in begin match get_variable0 mem ref with | ArgVar _ -> e - | ArgArg kn -> CTacRef (AbsKn kn) + | ArgArg kn -> Loc.tag ?loc @@ CTacRef (AbsKn kn) end -| CTacCst (loc, qid) -> +| CTacCst qid -> let knc = get_constructor () qid in - CTacCst (loc, AbsKn knc) -| CTacFun (loc, bnd, e) -> + Loc.tag ?loc @@ CTacCst (AbsKn knc) +| CTacFun (bnd, e) -> let fold (pats, accu) (pat, t) = let accu = ids_of_pattern accu pat in let pat = globalize_pattern ids pat in @@ -1224,12 +1211,12 @@ let rec globalize ids e = match e with let bnd, ids = List.fold_left fold ([], ids) bnd in let bnd = List.rev bnd in let e = globalize ids e in - CTacFun (loc, bnd, e) -| CTacApp (loc, e, el) -> + Loc.tag ?loc @@ CTacFun (bnd, e) +| CTacApp (e, el) -> let e = globalize ids e in let el = List.map (fun e -> globalize ids e) el in - CTacApp (loc, e, el) -| CTacLet (loc, isrec, bnd, e) -> + Loc.tag ?loc @@ CTacApp (e, el) +| CTacLet (isrec, bnd, e) -> let fold accu (pat, _, _) = ids_of_pattern accu pat in let ext = List.fold_left fold Id.Set.empty bnd in let eids = Id.Set.union ext ids in @@ -1239,48 +1226,48 @@ let rec globalize ids e = match e with (qid, t, globalize ids e) in let bnd = List.map map bnd in - CTacLet (loc, isrec, bnd, e) -| CTacCnv (loc, e, t) -> + Loc.tag ?loc @@ CTacLet (isrec, bnd, e) +| CTacCnv (e, t) -> let e = globalize ids e in - CTacCnv (loc, e, t) -| CTacSeq (loc, e1, e2) -> + Loc.tag ?loc @@ CTacCnv (e, t) +| CTacSeq (e1, e2) -> let e1 = globalize ids e1 in let e2 = globalize ids e2 in - CTacSeq (loc, e1, e2) -| CTacCse (loc, e, bl) -> + Loc.tag ?loc @@ CTacSeq (e1, e2) +| CTacCse (e, bl) -> let e = globalize ids e in let bl = List.map (fun b -> globalize_case ids b) bl in - CTacCse (loc, e, bl) -| CTacRec (loc, r) -> + Loc.tag ?loc @@ CTacCse (e, bl) +| CTacRec r -> let map (p, e) = let p = get_projection0 p in let e = globalize ids e in (AbsKn p, e) in - CTacRec (loc, List.map map r) -| CTacPrj (loc, e, p) -> + Loc.tag ?loc @@ CTacRec (List.map map r) +| CTacPrj (e, p) -> let e = globalize ids e in let p = get_projection0 p in - CTacPrj (loc, e, AbsKn p) -| CTacSet (loc, e, p, e') -> + Loc.tag ?loc @@ CTacPrj (e, AbsKn p) +| CTacSet (e, p, e') -> let e = globalize ids e in let p = get_projection0 p in let e' = globalize ids e' in - CTacSet (loc, e, AbsKn p, e') -| CTacExt (loc, tag, arg) -> + Loc.tag ?loc @@ CTacSet (e, AbsKn p, e') +| CTacExt (tag, arg) -> let arg = str (Tac2dyn.Arg.repr tag) in - CErrors.user_err ~loc (str "Cannot globalize generic arguments of type" ++ spc () ++ arg) + CErrors.user_err ?loc (str "Cannot globalize generic arguments of type" ++ spc () ++ arg) and globalize_case ids (p, e) = (globalize_pattern ids p, globalize ids e) -and globalize_pattern ids p = match p with +and globalize_pattern ids (loc, pr as p) = match pr with | CPatVar _ -> p -| CPatRef (loc, cst, pl) -> +| CPatRef (cst, pl) -> let knc = get_constructor () cst in let cst = AbsKn knc in let pl = List.map (fun p -> globalize_pattern ids p) pl in - CPatRef (loc, cst, pl) + Loc.tag ?loc @@ CPatRef (cst, pl) (** Kernel substitution *) @@ -1387,16 +1374,16 @@ let subst_or_relid subst ref = match ref with let kn' = subst_or_tuple subst_kn subst kn in if kn' == kn then ref else AbsKn kn' -let rec subst_rawtype subst t = match t with +let rec subst_rawtype subst (loc, tr as t) = match tr with | CTypVar _ -> t -| CTypArrow (loc, t1, t2) -> +| CTypArrow (t1, t2) -> let t1' = subst_rawtype subst t1 in let t2' = subst_rawtype subst t2 in - if t1' == t1 && t2' == t2 then t else CTypArrow (loc, t1', t2') -| CTypRef (loc, ref, tl) -> + if t1' == t1 && t2' == t2 then t else Loc.tag ?loc @@ CTypArrow (t1', t2') +| CTypRef (ref, tl) -> let ref' = subst_or_relid subst ref in let tl' = List.smartmap (fun t -> subst_rawtype subst t) tl in - if ref' == ref && tl' == tl then t else CTypRef (loc, ref', tl') + if ref' == ref && tl' == tl then t else Loc.tag ?loc @@ CTypRef (ref', tl') let subst_tacref subst ref = match ref with | RelId _ -> ref @@ -1413,35 +1400,35 @@ let subst_projection subst prj = match prj with let kn' = subst_kn subst kn in if kn' == kn then prj else AbsKn kn' -let rec subst_rawpattern subst p = match p with +let rec subst_rawpattern subst (loc, pr as p) = match pr with | CPatVar _ -> p -| CPatRef (loc, c, pl) -> +| CPatRef (c, pl) -> let pl' = List.smartmap (fun p -> subst_rawpattern subst p) pl in let c' = subst_or_relid subst c in - if pl' == pl && c' == c then p else CPatRef (loc, c', pl') + if pl' == pl && c' == c then p else Loc.tag ?loc @@ CPatRef (c', pl') (** Used for notations *) -let rec subst_rawexpr subst t = match t with +let rec subst_rawexpr subst (loc, tr as t) = match tr with | CTacAtm _ -> t | CTacRef ref -> let ref' = subst_tacref subst ref in - if ref' == ref then t else CTacRef ref' -| CTacCst (loc, ref) -> + if ref' == ref then t else Loc.tag ?loc @@ CTacRef ref' +| CTacCst ref -> let ref' = subst_or_relid subst ref in - if ref' == ref then t else CTacCst (loc, ref') -| CTacFun (loc, bnd, e) -> + if ref' == ref then t else Loc.tag ?loc @@ CTacCst ref' +| CTacFun (bnd, e) -> let map (na, t as p) = let t' = Option.smartmap (fun t -> subst_rawtype subst t) t in if t' == t then p else (na, t') in let bnd' = List.smartmap map bnd in let e' = subst_rawexpr subst e in - if bnd' == bnd && e' == e then t else CTacFun (loc, bnd', e') -| CTacApp (loc, e, el) -> + if bnd' == bnd && e' == e then t else Loc.tag ?loc @@ CTacFun (bnd', e') +| CTacApp (e, el) -> let e' = subst_rawexpr subst e in let el' = List.smartmap (fun e -> subst_rawexpr subst e) el in - if e' == e && el' == el then t else CTacApp (loc, e', el') -| CTacLet (loc, isrec, bnd, e) -> + if e' == e && el' == el then t else Loc.tag ?loc @@ CTacApp (e', el') +| CTacLet (isrec, bnd, e) -> let map (na, t, e as p) = let t' = Option.smartmap (fun t -> subst_rawtype subst t) t in let e' = subst_rawexpr subst e in @@ -1449,16 +1436,16 @@ let rec subst_rawexpr subst t = match t with in let bnd' = List.smartmap map bnd in let e' = subst_rawexpr subst e in - if bnd' == bnd && e' == e then t else CTacLet (loc, isrec, bnd', e') -| CTacCnv (loc, e, c) -> + if bnd' == bnd && e' == e then t else Loc.tag ?loc @@ CTacLet (isrec, bnd', e') +| CTacCnv (e, c) -> let e' = subst_rawexpr subst e in let c' = subst_rawtype subst c in - if c' == c && e' == e then t else CTacCnv (loc, e', c') -| CTacSeq (loc, e1, e2) -> + if c' == c && e' == e then t else Loc.tag ?loc @@ CTacCnv (e', c') +| CTacSeq (e1, e2) -> let e1' = subst_rawexpr subst e1 in let e2' = subst_rawexpr subst e2 in - if e1' == e1 && e2' == e2 then t else CTacSeq (loc, e1', e2') -| CTacCse (loc, e, bl) -> + if e1' == e1 && e2' == e2 then t else Loc.tag ?loc @@ CTacSeq (e1', e2') +| CTacCse (e, bl) -> let map (p, e as x) = let p' = subst_rawpattern subst p in let e' = subst_rawexpr subst e in @@ -1466,25 +1453,25 @@ let rec subst_rawexpr subst t = match t with in let e' = subst_rawexpr subst e in let bl' = List.smartmap map bl in - if e' == e && bl' == bl then t else CTacCse (loc, e', bl') -| CTacRec (loc, el) -> + if e' == e && bl' == bl then t else Loc.tag ?loc @@ CTacCse (e', bl') +| CTacRec el -> let map (prj, e as p) = let prj' = subst_projection subst prj in let e' = subst_rawexpr subst e in if prj' == prj && e' == e then p else (prj', e') in let el' = List.smartmap map el in - if el' == el then t else CTacRec (loc, el') -| CTacPrj (loc, e, prj) -> + if el' == el then t else Loc.tag ?loc @@ CTacRec el' +| CTacPrj (e, prj) -> let prj' = subst_projection subst prj in let e' = subst_rawexpr subst e in - if prj' == prj && e' == e then t else CTacPrj (loc, e', prj') -| CTacSet (loc, e, prj, r) -> + if prj' == prj && e' == e then t else Loc.tag ?loc @@ CTacPrj (e', prj') +| CTacSet (e, prj, r) -> let prj' = subst_projection subst prj in let e' = subst_rawexpr subst e in let r' = subst_rawexpr subst r in - if prj' == prj && e' == e && r' == r then t else CTacSet (loc, e', prj', r') -| CTacExt _ -> assert false (** Should not be generated by gloabalization *) + if prj' == prj && e' == e && r' == r then t else Loc.tag ?loc @@ CTacSet (e', prj', r') +| CTacExt _ -> assert false (** Should not be generated by globalization *) (** Registering *) diff --git a/src/tac2intern.mli b/src/tac2intern.mli index 95199d449d..045e657460 100644 --- a/src/tac2intern.mli +++ b/src/tac2intern.mli @@ -10,8 +10,8 @@ open Names open Mod_subst open Tac2expr -val loc_of_tacexpr : raw_tacexpr -> Loc.t -val loc_of_patexpr : raw_patexpr -> Loc.t +val loc_of_tacexpr : raw_tacexpr -> Loc.t option +val loc_of_patexpr : raw_patexpr -> Loc.t option val intern : raw_tacexpr -> glb_tacexpr * type_scheme val intern_typedef : (KerName.t * int) Id.Map.t -> raw_quant_typedef -> glb_quant_typedef @@ -41,8 +41,8 @@ val globalize : Id.Set.t -> raw_tacexpr -> raw_tacexpr (** Errors *) -val error_nargs_mismatch : Loc.t -> ltac_constructor -> int -> int -> 'a -val error_nparams_mismatch : Loc.t -> int -> int -> 'a +val error_nargs_mismatch : ?loc:Loc.t -> ltac_constructor -> int -> int -> 'a +val error_nparams_mismatch : ?loc:Loc.t -> int -> int -> 'a (** Misc *) diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 279ab53b67..063a52c738 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -25,10 +25,9 @@ let control_core n = kername control_prefix n let dummy_loc = Loc.make_loc (-1, -1) let constructor ?loc kn args = - let loc = Option.default dummy_loc loc in - let cst = CTacCst (loc, AbsKn (Other kn)) in + let cst = Loc.tag ?loc @@ CTacCst (AbsKn (Other kn)) in if List.is_empty args then cst - else CTacApp (loc, cst, args) + else Loc.tag ?loc @@ CTacApp (cst, args) let std_constructor ?loc name args = constructor ?loc (std_core name) args @@ -39,39 +38,35 @@ let std_proj ?loc name = let thunk e = let t_unit = coq_core "unit" in let loc = Tac2intern.loc_of_tacexpr e in - let var = [CPatVar (Some loc, Anonymous), Some (CTypRef (loc, AbsKn (Other t_unit), []))] in - CTacFun (loc, var, e) + let var = [Loc.tag ?loc @@ CPatVar (Anonymous), Some (Loc.tag ?loc @@ CTypRef (AbsKn (Other t_unit), []))] in + Loc.tag ?loc @@ CTacFun (var, e) let of_pair f g (loc, (e1, e2)) = - let loc = Option.default dummy_loc loc in - CTacApp (loc, CTacCst (loc, AbsKn (Tuple 2)), [f e1; g e2]) + Loc.tag ?loc @@ CTacApp (Loc.tag ?loc @@ CTacCst (AbsKn (Tuple 2)), [f e1; g e2]) let of_tuple ?loc el = match el with | [] -> - let loc = Option.default dummy_loc loc in - CTacCst (loc, AbsKn (Tuple 0)) + Loc.tag ?loc @@ CTacCst (AbsKn (Tuple 0)) | [e] -> e | el -> - let loc = Option.default dummy_loc loc in let len = List.length el in - CTacApp (loc, CTacCst (loc, AbsKn (Tuple len)), el) + Loc.tag ?loc @@ CTacApp (Loc.tag ?loc @@ CTacCst (AbsKn (Tuple len)), el) let of_int (loc, n) = - CTacAtm (Loc.tag ?loc (AtmInt n)) + Loc.tag ?loc @@ CTacAtm (AtmInt n) let of_option ?loc f opt = match opt with | None -> constructor ?loc (coq_core "None") [] | Some e -> constructor ?loc (coq_core "Some") [f e] let inj_wit ?loc wit x = - let loc = Option.default dummy_loc loc in - CTacExt (loc, wit, x) + Loc.tag ?loc @@ CTacExt (wit, x) let of_variable (loc, id) = let qid = Libnames.qualid_of_ident id in if Tac2env.is_constructor qid then CErrors.user_err ?loc (str "Invalid identifier") - else CTacRef (RelId (Loc.tag ?loc qid)) + else Loc.tag ?loc @@ CTacRef (RelId (Loc.tag ?loc qid)) let of_anti f = function | QExpr x -> f x @@ -171,10 +166,9 @@ let of_hyp_location ?loc ((occs, id), flag) = ] let of_clause (loc, cl) = - let loc = Option.default dummy_loc loc in - let hyps = of_option ~loc (fun l -> of_list ~loc of_hyp_location l) cl.q_onhyps in + let hyps = of_option ?loc (fun l -> of_list ?loc of_hyp_location l) cl.q_onhyps in let concl = of_occurrences cl.q_concl_occs in - CTacRec (loc, [ + Loc.tag ?loc @@ CTacRec ([ std_proj "on_hyps", hyps; std_proj "on_concl", concl; ]) @@ -191,8 +185,7 @@ let of_induction_clause (loc, cl) = let eqn = of_option ?loc of_intro_pattern_naming cl.indcl_eqn in let as_ = of_option ?loc of_or_and_intro_pattern cl.indcl_as in let in_ = of_option ?loc of_clause cl.indcl_in in - let loc = Option.default dummy_loc loc in - CTacRec (loc, [ + Loc.tag ?loc @@ CTacRec ([ std_proj "indcl_arg", arg; std_proj "indcl_eqn", eqn; std_proj "indcl_as", as_; @@ -216,36 +209,32 @@ let of_rewriting (loc, rew) = in let repeat = of_repeat rew.rew_repeat in let equatn = thunk (of_constr_with_bindings rew.rew_equatn) in - let loc = Option.default dummy_loc loc in - CTacRec (loc, [ + Loc.tag ?loc @@ CTacRec ([ std_proj "rew_orient", orient; std_proj "rew_repeat", repeat; std_proj "rew_equatn", equatn; ]) let of_hyp ?loc id = - let loc = Option.default dummy_loc loc in - let hyp = CTacRef (AbsKn (TacConstant (control_core "hyp"))) in - CTacApp (loc, hyp, [of_ident id]) + let hyp = Loc.tag ?loc @@ CTacRef (AbsKn (TacConstant (control_core "hyp"))) in + Loc.tag ?loc @@ CTacApp (hyp, [of_ident id]) let of_exact_hyp ?loc id = - let loc = Option.default dummy_loc loc in - let refine = CTacRef (AbsKn (TacConstant (control_core "refine"))) in - CTacApp (loc, refine, [thunk (of_hyp ~loc id)]) + let refine = Loc.tag ?loc @@ CTacRef (AbsKn (TacConstant (control_core "refine"))) in + Loc.tag ?loc @@ CTacApp (refine, [thunk (of_hyp ?loc id)]) let of_exact_var ?loc id = - let loc = Option.default dummy_loc loc in - let refine = CTacRef (AbsKn (TacConstant (control_core "refine"))) in - CTacApp (loc, refine, [thunk (of_variable id)]) + let refine = Loc.tag ?loc @@ CTacRef (AbsKn (TacConstant (control_core "refine"))) in + Loc.tag ?loc @@ CTacApp (refine, [thunk (of_variable id)]) let of_dispatch tacs = - let loc = Option.default dummy_loc (fst tacs) in + let (loc, _) = tacs in let default = function | Some e -> thunk e - | None -> thunk (CTacCst (loc, AbsKn (Tuple 0))) + | None -> thunk (Loc.tag ?loc @@ CTacCst (AbsKn (Tuple 0))) in - let map e = of_pair default (fun l -> of_list ~loc default l) (Loc.tag ~loc e) in - of_pair (fun l -> of_list ~loc default l) (fun r -> of_option ~loc map r) tacs + let map e = of_pair default (fun l -> of_list ?loc default l) (Loc.tag ?loc e) in + of_pair (fun l -> of_list ?loc default l) (fun r -> of_option ?loc map r) tacs let make_red_flag l = let open Genredexpr in @@ -287,14 +276,13 @@ let of_reference r = let of_strategy_flag (loc, flag) = let open Genredexpr in - let loc = Option.default dummy_loc loc in let flag = make_red_flag flag in - CTacRec (loc, [ - std_proj "rBeta", of_bool ~loc flag.rBeta; - std_proj "rMatch", of_bool ~loc flag.rMatch; - std_proj "rFix", of_bool ~loc flag.rFix; - std_proj "rCofix", of_bool ~loc flag.rCofix; - std_proj "rZeta", of_bool ~loc flag.rZeta; - std_proj "rDelta", of_bool ~loc flag.rDelta; - std_proj "rConst", of_list ~loc of_reference flag.rConst; + Loc.tag ?loc @@ CTacRec ([ + std_proj "rBeta", of_bool ?loc flag.rBeta; + std_proj "rMatch", of_bool ?loc flag.rMatch; + std_proj "rFix", of_bool ?loc flag.rFix; + std_proj "rCofix", of_bool ?loc flag.rCofix; + std_proj "rZeta", of_bool ?loc flag.rZeta; + std_proj "rDelta", of_bool ?loc flag.rDelta; + std_proj "rConst", of_list ?loc of_reference flag.rConst; ]) -- cgit v1.2.3 From a2302a48a96a6b635f5301f7cc6254acb58211bc Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 3 Sep 2017 19:39:47 +0200 Subject: Moving generic arguments to Tac2quote. --- src/g_ltac2.ml4 | 8 ++++---- src/tac2core.ml | 18 +++++++++--------- src/tac2env.ml | 7 ------- src/tac2env.mli | 15 --------------- src/tac2quote.ml | 18 ++++++++++++++---- src/tac2quote.mli | 18 ++++++++++++++++++ 6 files changed, 45 insertions(+), 39 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 338711e79c..fce4c9e159 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -87,10 +87,10 @@ let tac2mode = Gram.entry_create "vernac:ltac2_command" let ltac1_expr = (Obj.magic Pltac.tactic_expr : Tacexpr.raw_tactic_expr Gram.entry) let inj_wit wit loc x = Loc.tag ~loc @@ CTacExt (wit, x) -let inj_open_constr loc c = inj_wit Tac2env.wit_open_constr loc c -let inj_pattern loc c = inj_wit Tac2env.wit_pattern loc c -let inj_reference loc c = inj_wit Tac2env.wit_reference loc c -let inj_ltac1 loc e = inj_wit Tac2env.wit_ltac1 loc e +let inj_open_constr loc c = inj_wit Tac2quote.wit_open_constr loc c +let inj_pattern loc c = inj_wit Tac2quote.wit_pattern loc c +let inj_reference loc c = inj_wit Tac2quote.wit_reference loc c +let inj_ltac1 loc e = inj_wit Tac2quote.wit_ltac1 loc e let pattern_of_qualid ?loc id = if Tac2env.is_constructor (snd id) then Loc.tag ?loc @@ CPatRef (RelId id, []) diff --git a/src/tac2core.ml b/src/tac2core.ml index db8f989768..f5dd74d642 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -803,7 +803,7 @@ let () = ml_interp = interp; ml_print = print; } in - define_ml_object Tac2env.wit_constr obj + define_ml_object Tac2quote.wit_constr obj let () = let intern = intern_constr in @@ -815,7 +815,7 @@ let () = ml_interp = interp; ml_print = print; } in - define_ml_object Tac2env.wit_open_constr obj + define_ml_object Tac2quote.wit_open_constr obj let () = let interp _ id = return (ValExt (Value.val_ident, id)) in @@ -826,7 +826,7 @@ let () = ml_subst = (fun _ id -> id); ml_print = print; } in - define_ml_object Tac2env.wit_ident obj + define_ml_object Tac2quote.wit_ident obj let () = let intern self ist c = @@ -841,7 +841,7 @@ let () = ml_subst = Patternops.subst_pattern; ml_print = print; } in - define_ml_object Tac2env.wit_pattern obj + define_ml_object Tac2quote.wit_pattern obj let () = let intern self ist qid = match qid with @@ -867,7 +867,7 @@ let () = ml_interp = interp; ml_print = print; } in - define_ml_object Tac2env.wit_reference obj + define_ml_object Tac2quote.wit_reference obj let () = let intern self ist tac = @@ -892,7 +892,7 @@ let () = ml_interp = interp; ml_print = print; } in - define_ml_object Tac2env.wit_ltac1 obj + define_ml_object Tac2quote.wit_ltac1 obj (** Ltac2 in terms *) @@ -1070,9 +1070,9 @@ let () = add_expr_scope "dispatch" q_dispatch Tac2quote.of_dispatch let () = add_expr_scope "strategy" q_strategy_flag Tac2quote.of_strategy_flag let () = add_expr_scope "reference" q_reference Tac2quote.of_reference -let () = add_generic_scope "constr" Pcoq.Constr.constr wit_constr -let () = add_generic_scope "open_constr" Pcoq.Constr.constr wit_open_constr -let () = add_generic_scope "pattern" Pcoq.Constr.constr wit_pattern +let () = add_generic_scope "constr" Pcoq.Constr.constr Tac2quote.wit_constr +let () = add_generic_scope "open_constr" Pcoq.Constr.constr Tac2quote.wit_open_constr +let () = add_generic_scope "pattern" Pcoq.Constr.constr Tac2quote.wit_pattern (** seq scope, a bit hairy *) diff --git a/src/tac2env.ml b/src/tac2env.ml index 5a817df713..9aaaae0465 100644 --- a/src/tac2env.ml +++ b/src/tac2env.ml @@ -295,13 +295,6 @@ let std_prefix = let wit_ltac2 = Genarg.make0 "ltac2:value" -let wit_pattern = Arg.create "pattern" -let wit_reference = Arg.create "reference" -let wit_ident = Arg.create "ident" -let wit_constr = Arg.create "constr" -let wit_open_constr = Arg.create "open_constr" -let wit_ltac1 = Arg.create "ltac1" - let is_constructor qid = let (_, id) = repr_qualid qid in let id = Id.to_string id in diff --git a/src/tac2env.mli b/src/tac2env.mli index eb18dc8e39..e40958e1a0 100644 --- a/src/tac2env.mli +++ b/src/tac2env.mli @@ -133,21 +133,6 @@ val std_prefix : ModPath.t val wit_ltac2 : (raw_tacexpr, glb_tacexpr, Util.Empty.t) genarg_type -val wit_pattern : (Constrexpr.constr_expr, Pattern.constr_pattern) Arg.tag - -val wit_ident : (Id.t, Id.t) Arg.tag - -val wit_reference : (reference, Globnames.global_reference) Arg.tag -(** Beware, at the raw level, [Qualid [id]] has not the same meaning as - [Ident id]. The first is an unqualified global reference, the second is - the dynamic reference to id. *) - -val wit_constr : (Constrexpr.constr_expr, Glob_term.glob_constr) Arg.tag - -val wit_open_constr : (Constrexpr.constr_expr, Glob_term.glob_constr) Arg.tag - -val wit_ltac1 : (Tacexpr.raw_tactic_expr, Tacexpr.glob_tactic_expr) Arg.tag - (** {5 Helper functions} *) val is_constructor : qualid -> bool diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 063a52c738..132716c37e 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -9,9 +9,19 @@ open Pp open Names open Util +open Tac2dyn open Tac2expr open Tac2qexpr +(** Generic arguments *) + +let wit_pattern = Arg.create "pattern" +let wit_reference = Arg.create "reference" +let wit_ident = Arg.create "ident" +let wit_constr = Arg.create "constr" +let wit_open_constr = Arg.create "open_constr" +let wit_ltac1 = Arg.create "ltac1" + (** Syntactic quoting of expressions. *) let control_prefix = @@ -72,15 +82,15 @@ let of_anti f = function | QExpr x -> f x | QAnti id -> of_variable id -let of_ident (loc, id) = inj_wit ?loc Tac2env.wit_ident id +let of_ident (loc, id) = inj_wit ?loc wit_ident id let of_constr c = let loc = Constrexpr_ops.constr_loc c in - inj_wit ?loc Tac2env.wit_constr c + inj_wit ?loc wit_constr c let of_open_constr c = let loc = Constrexpr_ops.constr_loc c in - inj_wit ?loc Tac2env.wit_open_constr c + inj_wit ?loc wit_open_constr c let of_bool ?loc b = let c = if b then coq_core "true" else coq_core "false" in @@ -270,7 +280,7 @@ let make_red_flag l = let of_reference r = let of_ref ref = let loc = Libnames.loc_of_reference ref in - inj_wit ?loc Tac2env.wit_reference ref + inj_wit ?loc wit_reference ref in of_anti of_ref r diff --git a/src/tac2quote.mli b/src/tac2quote.mli index b2687f01a3..440759ada7 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -8,6 +8,7 @@ open Loc open Names +open Tac2dyn open Tac2qexpr open Tac2expr @@ -64,3 +65,20 @@ val of_exact_var : ?loc:Loc.t -> Id.t located -> raw_tacexpr val of_dispatch : dispatch -> raw_tacexpr val of_strategy_flag : strategy_flag -> raw_tacexpr + +(** {5 Generic arguments} *) + +val wit_pattern : (Constrexpr.constr_expr, Pattern.constr_pattern) Arg.tag + +val wit_ident : (Id.t, Id.t) Arg.tag + +val wit_reference : (Libnames.reference, Globnames.global_reference) Arg.tag +(** Beware, at the raw level, [Qualid [id]] has not the same meaning as + [Ident id]. The first is an unqualified global reference, the second is + the dynamic reference to id. *) + +val wit_constr : (Constrexpr.constr_expr, Glob_term.glob_constr) Arg.tag + +val wit_open_constr : (Constrexpr.constr_expr, Glob_term.glob_constr) Arg.tag + +val wit_ltac1 : (Tacexpr.raw_tactic_expr, Tacexpr.glob_tactic_expr) Arg.tag -- cgit v1.2.3 From 65daf8fca747d385b2985e4e5e91894738f6fcf1 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 3 Sep 2017 19:45:23 +0200 Subject: Introducing a macro for constr matching. --- src/g_ltac2.ml4 | 21 ++++++++++++++- src/tac2core.ml | 2 ++ src/tac2entries.ml | 1 + src/tac2entries.mli | 1 + src/tac2qexpr.mli | 8 ++++++ src/tac2quote.ml | 74 +++++++++++++++++++++++++++++++++++++++++++++++++---- src/tac2quote.mli | 3 +++ theories/Pattern.v | 5 ++++ 8 files changed, 109 insertions(+), 6 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index fce4c9e159..5d5bc6b941 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -368,7 +368,7 @@ let loc_of_ne_list l = Loc.merge_opt (fst (List.hd l)) (fst (List.last l)) GEXTEND Gram GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause q_rewriting q_clause q_dispatch q_occurrences q_strategy_flag - q_reference q_with_bindings; + q_reference q_with_bindings q_constr_matching; anti: [ [ "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] ; @@ -661,6 +661,25 @@ GEXTEND Gram q_strategy_flag: [ [ flag = strategy_flag -> flag ] ] ; + match_pattern: + [ [ IDENT "context"; id = OPT Prim.ident; + "["; pat = Constr.lconstr_pattern; "]" -> (Some id, pat) + | pat = Constr.lconstr_pattern -> (None, pat) ] ] + ; + match_rule: + [ [ mp = match_pattern; "=>"; tac = tac2expr -> + match mp with + | None, pat -> Loc.tag ~loc:!@loc @@ QConstrMatchPattern (pat, tac) + | Some oid, pat -> Loc.tag ~loc:!@loc @@ QConstrMatchContext (oid, pat, tac) + ] ] + ; + match_list: + [ [ mrl = LIST1 match_rule SEP "|" -> Loc.tag ~loc:!@loc @@ mrl + | "|"; mrl = LIST1 match_rule SEP "|" -> Loc.tag ~loc:!@loc @@ mrl ] ] + ; + q_constr_matching: + [ [ m = match_list -> m ] ] + ; END (** Extension of constr syntax *) diff --git a/src/tac2core.ml b/src/tac2core.ml index f5dd74d642..39fcff0c73 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -7,6 +7,7 @@ (************************************************************************) open CSig +open Util open Pp open Names open Genarg @@ -1069,6 +1070,7 @@ let () = add_expr_scope "occurrences" q_occurrences Tac2quote.of_occurrences let () = add_expr_scope "dispatch" q_dispatch Tac2quote.of_dispatch let () = add_expr_scope "strategy" q_strategy_flag Tac2quote.of_strategy_flag let () = add_expr_scope "reference" q_reference Tac2quote.of_reference +let () = add_expr_scope "constr_matching" q_constr_matching Tac2quote.of_constr_matching let () = add_generic_scope "constr" Pcoq.Constr.constr Tac2quote.wit_constr let () = add_generic_scope "open_constr" Pcoq.Constr.constr Tac2quote.wit_open_constr diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 9c5d9a659b..34022b86c9 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -36,6 +36,7 @@ let q_dispatch = Pcoq.Gram.entry_create "tactic:q_dispatch" let q_occurrences = Pcoq.Gram.entry_create "tactic:q_occurrences" let q_reference = Pcoq.Gram.entry_create "tactic:q_reference" let q_strategy_flag = Pcoq.Gram.entry_create "tactic:q_strategy_flag" +let q_constr_matching = Pcoq.Gram.entry_create "tactic:q_constr_matching" end (** Tactic definition *) diff --git a/src/tac2entries.mli b/src/tac2entries.mli index 7ed45e62e5..dde877666a 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -71,6 +71,7 @@ val q_dispatch : dispatch Pcoq.Gram.entry val q_occurrences : occurrences Pcoq.Gram.entry val q_reference : Libnames.reference or_anti Pcoq.Gram.entry val q_strategy_flag : strategy_flag Pcoq.Gram.entry +val q_constr_matching : constr_matching Pcoq.Gram.entry end (** {5 Hooks} *) diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index 7c774a5c80..a284c1e756 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -115,3 +115,11 @@ type red_flag_r = type red_flag = red_flag_r located type strategy_flag = red_flag list located + +type constr_match_branch_r = +| QConstrMatchPattern of Constrexpr.constr_expr * raw_tacexpr +| QConstrMatchContext of Id.t option * Constrexpr.constr_expr * raw_tacexpr + +type constr_match_branch = constr_match_branch_r located + +type constr_matching = constr_match_branch list located diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 132716c37e..d38d7414fe 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -24,13 +24,21 @@ let wit_ltac1 = Arg.create "ltac1" (** Syntactic quoting of expressions. *) -let control_prefix = - MPfile (DirPath.make (List.map Id.of_string ["Control"; "Ltac2"])) +let prefix_gen n = + MPfile (DirPath.make (List.map Id.of_string [n; "Ltac2"])) + +let control_prefix = prefix_gen "Control" +let pattern_prefix = prefix_gen "Pattern" +let array_prefix = prefix_gen "Array" let kername prefix n = KerName.make2 prefix (Label.of_id (Id.of_string_soft n)) let std_core n = kername Tac2env.std_prefix n let coq_core n = kername Tac2env.coq_prefix n let control_core n = kername control_prefix n +let pattern_core n = kername pattern_prefix n + +let global_ref ?loc kn = + Loc.tag ?loc @@ CTacRef (AbsKn (TacConstant kn)) let dummy_loc = Loc.make_loc (-1, -1) @@ -226,15 +234,15 @@ let of_rewriting (loc, rew) = ]) let of_hyp ?loc id = - let hyp = Loc.tag ?loc @@ CTacRef (AbsKn (TacConstant (control_core "hyp"))) in + let hyp = global_ref ?loc (control_core "hyp") in Loc.tag ?loc @@ CTacApp (hyp, [of_ident id]) let of_exact_hyp ?loc id = - let refine = Loc.tag ?loc @@ CTacRef (AbsKn (TacConstant (control_core "refine"))) in + let refine = global_ref ?loc (control_core "refine") in Loc.tag ?loc @@ CTacApp (refine, [thunk (of_hyp ?loc id)]) let of_exact_var ?loc id = - let refine = Loc.tag ?loc @@ CTacRef (AbsKn (TacConstant (control_core "refine"))) in + let refine = global_ref ?loc (control_core "refine") in Loc.tag ?loc @@ CTacApp (refine, [thunk (of_variable id)]) let of_dispatch tacs = @@ -296,3 +304,59 @@ let of_strategy_flag (loc, flag) = std_proj "rDelta", of_bool ?loc flag.rDelta; std_proj "rConst", of_list ?loc of_reference flag.rConst; ]) + +let pattern_vars pat = + let rec aux () accu pat = match pat.CAst.v with + | Constrexpr.CPatVar id -> Id.Set.add id accu + | Constrexpr.CEvar (id, []) -> Id.Set.add id accu + | _ -> + Topconstr.fold_constr_expr_with_binders (fun _ () -> ()) aux () accu pat + in + aux () Id.Set.empty pat + +let of_constr_matching (loc, m) = + let check_id loc id = + if Tac2env.is_constructor (Libnames.qualid_of_ident id) then + CErrors.user_err ?loc (str "Invalid pattern binding name " ++ Id.print id) + in + let abstract_vars loc pat tac = + let vars = pattern_vars pat in + let na, tac = + if Id.Set.is_empty vars then (Anonymous, tac) + else + (** Trick: in order not to shadow a variable nor to choose an arbitrary + name, we reuse one which is going to be shadowed by the matched + variables anyways. *) + let id0 = Id.Set.choose vars in + let build_bindings id (n, accu) = + let () = check_id loc id in + let get = global_ref ?loc (kername array_prefix "get") in + let args = [of_variable (loc, id0); of_int (loc, n)] in + let e = Loc.tag ?loc @@ CTacApp (get, args) in + let accu = (Loc.tag ?loc @@ CPatVar (Name id), None, e) :: accu in + (n + 1, accu) + in + let (_, bnd) = Id.Set.fold build_bindings vars (0, []) in + let tac = Loc.tag ?loc @@ CTacLet (false, bnd, tac) in + (Name id0, tac) + in + Loc.tag ?loc @@ CTacFun ([Loc.tag ?loc @@ CPatVar na, None], tac) + in + let map (loc, p) = match p with + | QConstrMatchPattern (pat, tac) -> + let e = abstract_vars loc pat tac in + let pat = inj_wit ?loc wit_pattern pat in + constructor ?loc (pattern_core "ConstrMatchPattern") [pat; e] + | QConstrMatchContext (id, pat, tac) -> + let e = abstract_vars loc pat tac in + let na = match id with + | None -> Anonymous + | Some id -> + let () = check_id loc id in + Name id + in + let e = Loc.tag ?loc @@ CTacFun ([Loc.tag ?loc @@ CPatVar na, None], e) in + let pat = inj_wit ?loc wit_pattern pat in + constructor ?loc (pattern_core "ConstrMatchContext") [pat; e] + in + of_list ?loc map m diff --git a/src/tac2quote.mli b/src/tac2quote.mli index 440759ada7..c3374ac24e 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open Util open Loc open Names open Tac2dyn @@ -66,6 +67,8 @@ val of_dispatch : dispatch -> raw_tacexpr val of_strategy_flag : strategy_flag -> raw_tacexpr +val of_constr_matching : constr_matching -> raw_tacexpr + (** {5 Generic arguments} *) val wit_pattern : (Constrexpr.constr_expr, Pattern.constr_pattern) Arg.tag diff --git a/theories/Pattern.v b/theories/Pattern.v index c2ba4162e8..ab3135f189 100644 --- a/theories/Pattern.v +++ b/theories/Pattern.v @@ -12,6 +12,11 @@ Ltac2 Type t := pattern. Ltac2 Type context. +Ltac2 Type 'a constr_match := [ +| ConstrMatchPattern (pattern, constr array -> 'a) +| ConstrMatchContext (pattern, constr -> constr array -> 'a) +]. + Ltac2 @ external matches : t -> constr -> (ident * constr) list := "ltac2" "pattern_matches". (** If the term matches the pattern, returns the bound variables. If it doesn't, -- cgit v1.2.3 From 34912844e18ef84d88af87e1dca05ab0426871c9 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 4 Sep 2017 00:02:06 +0200 Subject: Proper anomalies for missing registered primitives. --- src/tac2intern.ml | 12 ++++++++++-- 1 file changed, 10 insertions(+), 2 deletions(-) diff --git a/src/tac2intern.ml b/src/tac2intern.ml index c1fd281808..1dba57c4c1 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -651,10 +651,18 @@ let rec intern_rec env (loc, e) = match e with let sch = Id.Map.find id env.env_var in (GTacVar id, fresh_mix_type_scheme env sch) | ArgArg (TacConstant kn) -> - let { Tac2env.gdata_type = sch }, _ = Tac2env.interp_global kn in + let { Tac2env.gdata_type = sch }, _ = + try Tac2env.interp_global kn + with Not_found -> + CErrors.anomaly (str "Missing hardwired primitive " ++ KerName.print kn) + in (GTacRef kn, fresh_type_scheme env sch) | ArgArg (TacAlias kn) -> - let e = Tac2env.interp_alias kn in + let e = + try Tac2env.interp_alias kn + with Not_found -> + CErrors.anomaly (str "Missing hardwired alias " ++ KerName.print kn) + in intern_rec env e end | CTacCst qid -> -- cgit v1.2.3 From 102cfe76bc42d3139c79eca59eb782fcf644317b Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 3 Sep 2017 23:52:12 +0200 Subject: Implementing lazy matching over terms. --- src/tac2core.ml | 36 ++++++++++++++++++++++++++++++++++++ tests/matching.v | 17 +++++++++++++++++ theories/Notations.v | 31 ++++++++++++++++++++++++++++++- theories/Pattern.v | 10 +++++++++- 4 files changed, 92 insertions(+), 2 deletions(-) create mode 100644 tests/matching.v diff --git a/src/tac2core.ml b/src/tac2core.ml index 39fcff0c73..bbf95c7056 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -544,6 +544,42 @@ let () = define2 "pattern_matches_subterm" begin fun pat c -> end end +let () = define2 "pattern_matches_vect" begin fun pat c -> + let pat = Value.to_pattern pat in + let c = Value.to_constr c in + pf_apply begin fun env sigma -> + let ans = + try Some (Constr_matching.matches env sigma pat c) + with Constr_matching.PatternMatchingFailure -> None + in + begin match ans with + | None -> Proofview.tclZERO err_matchfailure + | Some ans -> + let ans = Id.Map.bindings ans in + let ans = Array.map_of_list snd ans in + return (Value.of_array Value.of_constr ans) + end + end +end + +let () = define2 "pattern_matches_subterm_vect" begin fun pat c -> + let pat = Value.to_pattern pat in + let c = Value.to_constr c in + let open Constr_matching in + let rec of_ans s = match IStream.peek s with + | IStream.Nil -> Proofview.tclZERO err_matchfailure + | IStream.Cons ({ m_sub = (_, sub); m_ctx }, s) -> + let ans = Id.Map.bindings sub in + let ans = Array.map_of_list snd ans in + let ans = Value.of_tuple [| Value.of_constr m_ctx; Value.of_array Value.of_constr ans |] in + Proofview.tclOR (return ans) (fun _ -> of_ans s) + in + pf_apply begin fun env sigma -> + let ans = Constr_matching.match_appsubterm env sigma pat c in + of_ans ans + end +end + let () = define2 "pattern_instantiate" begin fun ctx c -> let ctx = EConstr.Unsafe.to_constr (Value.to_constr ctx) in let c = EConstr.Unsafe.to_constr (Value.to_constr c) in diff --git a/tests/matching.v b/tests/matching.v new file mode 100644 index 0000000000..d21c505cdf --- /dev/null +++ b/tests/matching.v @@ -0,0 +1,17 @@ +Require Import Ltac2.Ltac2 Ltac2.Notations. + +Goal True -> False. +Proof. +Fail +let b := { contents := true } in +let f c := + match b.(contents) with + | true => Message.print (Message.of_constr c); b.(contents) := false; fail + | false => () + end +in +(** This fails because the matching is not allowed to backtrack once + it commits to a branch*) +lazy_match! '(nat -> bool) with context [?a] => f a end. +lazy_match! Control.goal () with ?a -> ?b => Message.print (Message.of_constr b) end. +Abort. diff --git a/theories/Notations.v b/theories/Notations.v index 411367eab1..93c9dd1798 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -7,7 +7,36 @@ (************************************************************************) Require Import Ltac2.Init. -Require Ltac2.Control Ltac2.Int Ltac2.Std. +Require Ltac2.Control Ltac2.Pattern Ltac2.Array Ltac2.Int Ltac2.Std. + +(** Constr matching *) + +Ltac2 lazy_match0 t pats := + let rec interp m := match m with + | [] => Control.zero Match_failure + | p :: m => + match p with + | Pattern.ConstrMatchPattern pat f => + Control.plus + (fun _ => + let bind := Pattern.matches_vect pat t in + fun _ => f bind + ) + (fun _ => interp m) + | Pattern.ConstrMatchContext pat f => + Control.plus + (fun _ => + let ((context, bind)) := Pattern.matches_subterm_vect pat t in + fun _ => f context bind + ) + (fun _ => interp m) + end + end in + let ans := Control.once (fun () => interp pats) in + ans (). + +Ltac2 Notation "lazy_match!" t(tactic(6)) "with" m(constr_matching) "end" := + lazy_match0 t m. (** Tacticals *) diff --git a/theories/Pattern.v b/theories/Pattern.v index ab3135f189..a672ad0fe7 100644 --- a/theories/Pattern.v +++ b/theories/Pattern.v @@ -14,7 +14,7 @@ Ltac2 Type context. Ltac2 Type 'a constr_match := [ | ConstrMatchPattern (pattern, constr array -> 'a) -| ConstrMatchContext (pattern, constr -> constr array -> 'a) +| ConstrMatchContext (pattern, context -> constr array -> 'a) ]. Ltac2 @ external matches : t -> constr -> (ident * constr) list := @@ -30,6 +30,14 @@ Ltac2 @ external matches_subterm : t -> constr -> context * ((ident * constr) li value compared to [matches] is the context of the match, to be filled with the instantiate function. *) +Ltac2 @ external matches_vect : t -> constr -> constr array := + "ltac2" "pattern_matches_vect". +(** Internal version of [matches] that does not return the identifiers. *) + +Ltac2 @ external matches_subterm_vect : t -> constr -> context * constr array := + "ltac2" "pattern_matches_subterm_vect". +(** Internal version of [matches_subterms] that does not return the identifiers. *) + Ltac2 @ external instantiate : context -> constr -> constr := "ltac2" "pattern_instantiate". (** Fill the hole of a context with the given term. *) -- cgit v1.2.3 From e634eb23010a3dee3fddcdd3d7d5588c3b40e1f6 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 4 Sep 2017 13:39:17 +0200 Subject: Implementing multi-match and simple match over constrs. --- tests/matching.v | 10 ++++++++++ theories/Notations.v | 31 +++++++++++++++++++++++++++++++ 2 files changed, 41 insertions(+) diff --git a/tests/matching.v b/tests/matching.v index d21c505cdf..f43e0121ef 100644 --- a/tests/matching.v +++ b/tests/matching.v @@ -14,4 +14,14 @@ in it commits to a branch*) lazy_match! '(nat -> bool) with context [?a] => f a end. lazy_match! Control.goal () with ?a -> ?b => Message.print (Message.of_constr b) end. + +(** This one works by taking the second match context, i.e. ?a := nat *) +let b := { contents := true } in +let f c := + match b.(contents) with + | true => b.(contents) := false; fail + | false => Message.print (Message.of_constr c) + end +in +match! '(nat -> bool) with context [?a] => f a end. Abort. diff --git a/theories/Notations.v b/theories/Notations.v index 93c9dd1798..5ed47336ad 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -38,6 +38,37 @@ Ltac2 lazy_match0 t pats := Ltac2 Notation "lazy_match!" t(tactic(6)) "with" m(constr_matching) "end" := lazy_match0 t m. +Ltac2 multi_match0 t pats := + let rec interp m := match m with + | [] => Control.zero Match_failure + | p :: m => + match p with + | Pattern.ConstrMatchPattern pat f => + Control.plus + (fun _ => + let bind := Pattern.matches_vect pat t in + f bind + ) + (fun _ => interp m) + | Pattern.ConstrMatchContext pat f => + Control.plus + (fun _ => + let ((context, bind)) := Pattern.matches_subterm_vect pat t in + f context bind + ) + (fun _ => interp m) + end + end in + interp pats. + +Ltac2 Notation "multi_match!" t(tactic(6)) "with" m(constr_matching) "end" := + multi_match0 t m. + +Ltac2 one_match0 t m := Control.once (fun _ => multi_match0 t m). + +Ltac2 Notation "match!" t(tactic(6)) "with" m(constr_matching) "end" := + one_match0 t m. + (** Tacticals *) Ltac2 orelse t f := -- cgit v1.2.3 From 818c49240f2ee6fccd38a556c7e90126606e1837 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 4 Sep 2017 14:45:32 +0200 Subject: Fix coq/ltac2#17: Assertion failed. on only shelved goals remaining. We check that the goal tactic is focussed before calling enter_one. --- src/tac2core.ml | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/src/tac2core.ml b/src/tac2core.ml index bbf95c7056..f4486bf0c8 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -116,6 +116,12 @@ let wrap f = let wrap_unit f = return () >>= fun () -> f (); return v_unit +let assert_focussed = + Proofview.Goal.goals >>= fun gls -> + match gls with + | [_] -> Proofview.tclUNIT () + | [] | _ :: _ :: _ -> throw err_notfocussed + let pf_apply f = Proofview.Goal.goals >>= function | [] -> @@ -682,6 +688,7 @@ end (** unit -> constr *) let () = define0 "goal" begin + assert_focussed >>= fun () -> Proofview.Goal.enter_one begin fun gl -> let concl = Tacmach.New.pf_nf_concl gl in return (Value.of_constr concl) -- cgit v1.2.3 From cd3819db675bd42510eac1bd616ca20e33e7d997 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 1 Sep 2017 00:57:33 +0200 Subject: Closures now wear the constant they originated from. --- src/tac2core.ml | 1 + src/tac2entries.ml | 2 +- src/tac2env.ml | 14 ++++++++------ src/tac2expr.mli | 2 ++ src/tac2interp.ml | 8 ++++---- 5 files changed, 16 insertions(+), 11 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index f4486bf0c8..793cb3e535 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -647,6 +647,7 @@ let () = define1 "case" begin fun f -> Proofview.tclCASE (thaw f) >>= begin function | Proofview.Next (x, k) -> let k = { + clos_ref = None; clos_env = Id.Map.singleton k_var (Value.of_ext Value.val_kont k); clos_var = [Name e_var]; clos_exp = GTacPrm (prm_apply_kont_h, [GTacVar k_var; GTacVar e_var]); diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 34022b86c9..0754108505 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -792,7 +792,7 @@ let call ~default e = let loc = loc_of_tacexpr e in let (e, t) = intern e in let () = check_unit ?loc t in - let tac = Tac2interp.interp Id.Map.empty e in + let tac = Tac2interp.interp Tac2interp.empty_environment e in solve default (Proofview.tclIGNORE tac) (** Primitive algebraic types than can't be defined Coq-side *) diff --git a/src/tac2env.ml b/src/tac2env.ml index 9aaaae0465..c04eaf7b0c 100644 --- a/src/tac2env.ml +++ b/src/tac2env.ml @@ -53,30 +53,32 @@ let empty_state = { let ltac_state = Summary.ref empty_state ~name:"ltac2-state" (** Get a dynamic value from syntactical value *) -let rec eval_pure = function +let rec eval_pure kn = function | GTacAtm (AtmInt n) -> ValInt n | GTacRef kn -> let { gdata_expr = e } = try KNmap.find kn ltac_state.contents.ltac_tactics with Not_found -> assert false in - eval_pure e + eval_pure (Some kn) e | GTacFun (na, e) -> - ValCls { clos_env = Id.Map.empty; clos_var = na; clos_exp = e } + ValCls { clos_ref = kn; clos_env = Id.Map.empty; clos_var = na; clos_exp = e } | GTacCst (_, n, []) -> ValInt n -| GTacCst (_, n, el) -> ValBlk (n, Array.map_of_list eval_pure el) -| GTacOpn (kn, el) -> ValOpn (kn, Array.map_of_list eval_pure el) +| GTacCst (_, n, el) -> ValBlk (n, Array.map_of_list eval_unnamed el) +| GTacOpn (kn, el) -> ValOpn (kn, Array.map_of_list eval_unnamed el) | GTacAtm (AtmStr _) | GTacLet _ | GTacVar _ | GTacSet _ | GTacApp _ | GTacCse _ | GTacPrj _ | GTacPrm _ | GTacExt _ | GTacWth _ -> anomaly (Pp.str "Term is not a syntactical value") +and eval_unnamed e = eval_pure None e + let define_global kn e = let state = !ltac_state in ltac_state := { state with ltac_tactics = KNmap.add kn e state.ltac_tactics } let interp_global kn = let data = KNmap.find kn ltac_state.contents.ltac_tactics in - (data, eval_pure data.gdata_expr) + (data, eval_pure (Some kn) data.gdata_expr) let define_constructor kn t = let state = !ltac_state in diff --git a/src/tac2expr.mli b/src/tac2expr.mli index 1045063cd2..470323e7c7 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -201,6 +201,8 @@ and closure = { (** Bound variables *) clos_exp : glb_tacexpr; (** Body *) + clos_ref : ltac_constant option; + (** Global constant from which the closure originates *) } type ml_tactic = valexpr list -> valexpr Proofview.tactic diff --git a/src/tac2interp.ml b/src/tac2interp.ml index c15331571b..7bcfad1be1 100644 --- a/src/tac2interp.ml +++ b/src/tac2interp.ml @@ -47,7 +47,7 @@ let rec interp ist = function | GTacVar id -> return (get_var ist id) | GTacRef qid -> return (get_ref ist qid) | GTacFun (ids, e) -> - let cls = { clos_env = ist; clos_var = ids; clos_exp = e } in + let cls = { clos_ref = None; clos_env = ist; clos_var = ids; clos_exp = e } in return (ValCls cls) | GTacApp (f, args) -> interp ist f >>= fun f -> @@ -63,7 +63,7 @@ let rec interp ist = function | GTacLet (true, el, e) -> let map (na, e) = match e with | GTacFun (ids, e) -> - let cls = { clos_env = ist; clos_var = ids; clos_exp = e } in + let cls = { clos_ref = None; clos_env = ist; clos_var = ids; clos_exp = e } in na, cls | _ -> anomaly (str "Ill-formed recursive function") in @@ -102,12 +102,12 @@ let rec interp ist = function tpe.Tac2env.ml_interp ist e and interp_app f args = match f with -| ValCls { clos_env = ist; clos_var = ids; clos_exp = e } -> +| ValCls { clos_env = ist; clos_var = ids; clos_exp = e; clos_ref = kn } -> let rec push ist ids args = match ids, args with | [], [] -> interp ist e | [], _ :: _ -> interp ist e >>= fun f -> interp_app f args | _ :: _, [] -> - let cls = { clos_env = ist; clos_var = ids; clos_exp = e } in + let cls = { clos_ref = kn; clos_env = ist; clos_var = ids; clos_exp = e } in return (ValCls cls) | id :: ids, arg :: args -> push (push_name ist id arg) ids args in -- cgit v1.2.3 From dd2a9aa0fd0a8d725f131223a4e0a01de8a98e1e Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 4 Sep 2017 15:41:50 +0200 Subject: Updated gitignore. --- .gitignore | 1 + 1 file changed, 1 insertion(+) diff --git a/.gitignore b/.gitignore index ffdea1320c..50ed772be3 100644 --- a/.gitignore +++ b/.gitignore @@ -12,3 +12,4 @@ Makefile.coq.conf *.o *.a *.aux +tests/*.log -- cgit v1.2.3 From d80e854d6827252676c2c504bb3108152a94d629 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 4 Sep 2017 15:24:33 +0200 Subject: Quick-and-dirty backtrace mechanism for the interpreter. --- doc/ltac2.md | 5 ++ src/tac2core.ml | 254 +++++++++++++++++++++++++++------------------------- src/tac2core.mli | 2 +- src/tac2entries.ml | 35 ++++++++ src/tac2expr.mli | 14 ++- src/tac2ffi.ml | 5 +- src/tac2interp.ml | 43 ++++----- src/tac2interp.mli | 6 +- src/tac2stdlib.ml | 190 ++++++++++++++++++++------------------- src/tac2tactics.ml | 48 +++++----- src/tac2tactics.mli | 23 ++--- 11 files changed, 340 insertions(+), 285 deletions(-) diff --git a/doc/ltac2.md b/doc/ltac2.md index dd0dc391c6..c1216d8f89 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -703,6 +703,11 @@ foo 0 ↦ (fun x => x ()) (fun _ => 0) Note that abbreviations are not typechecked at all, and may result in typing errors after expansion. +# Debug + +When the option `Ltac2 Backtrace` is set, toplevel failures will be printed with +a backtrace. + # Compatibility layer with Ltac1 ## Ltac1 from Ltac2 diff --git a/src/tac2core.ml b/src/tac2core.ml index 793cb3e535..17fa7c28f4 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -14,7 +14,6 @@ open Genarg open Tac2env open Tac2dyn open Tac2expr -open Tac2interp open Tac2entries.Pltac open Proofview.Notations @@ -90,22 +89,26 @@ let of_result f = function (** Stdlib exceptions *) -let err_notfocussed = - LtacError (coq_core "Not_focussed", [||]) +let err_notfocussed bt = + Tac2interp.LtacError (coq_core "Not_focussed", [||], bt) -let err_outofbounds = - LtacError (coq_core "Out_of_bounds", [||]) +let err_outofbounds bt = + Tac2interp.LtacError (coq_core "Out_of_bounds", [||], bt) -let err_notfound = - LtacError (coq_core "Not_found", [||]) +let err_notfound bt = + Tac2interp.LtacError (coq_core "Not_found", [||], bt) -let err_matchfailure = - LtacError (coq_core "Match_failure", [||]) +let err_matchfailure bt = + Tac2interp.LtacError (coq_core "Match_failure", [||], bt) (** Helper functions *) -let thaw f = interp_app f [v_unit] -let throw e = Proofview.tclLIFT (Proofview.NonLogical.raise e) +let thaw bt f = Tac2interp.interp_app bt f [v_unit] +let throw bt e = Proofview.tclLIFT (Proofview.NonLogical.raise (e bt)) + +let set_bt bt e = match e with +| Tac2interp.LtacError (kn, args, _) -> Tac2interp.LtacError (kn, args, bt) +| e -> e let return x = Proofview.tclUNIT x let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } @@ -116,13 +119,13 @@ let wrap f = let wrap_unit f = return () >>= fun () -> f (); return v_unit -let assert_focussed = +let assert_focussed bt = Proofview.Goal.goals >>= fun gls -> match gls with | [_] -> Proofview.tclUNIT () - | [] | _ :: _ :: _ -> throw err_notfocussed + | [] | _ :: _ :: _ -> throw bt err_notfocussed -let pf_apply f = +let pf_apply bt f = Proofview.Goal.goals >>= function | [] -> Proofview.tclENV >>= fun env -> @@ -132,61 +135,61 @@ let pf_apply f = gl >>= fun gl -> f (Proofview.Goal.env gl) (Tacmach.New.project gl) | _ :: _ :: _ -> - throw err_notfocussed + throw bt err_notfocussed (** Primitives *) -let define0 name f = Tac2env.define_primitive (pname name) begin function -| [_] -> f +let define0 name f = Tac2env.define_primitive (pname name) begin fun bt arg -> match arg with +| [_] -> f bt | _ -> assert false end -let define1 name f = Tac2env.define_primitive (pname name) begin function -| [x] -> f x +let define1 name f = Tac2env.define_primitive (pname name) begin fun bt arg -> match arg with +| [x] -> f bt x | _ -> assert false end -let define2 name f = Tac2env.define_primitive (pname name) begin function -| [x; y] -> f x y +let define2 name f = Tac2env.define_primitive (pname name) begin fun bt arg -> match arg with +| [x; y] -> f bt x y | _ -> assert false end -let define3 name f = Tac2env.define_primitive (pname name) begin function -| [x; y; z] -> f x y z +let define3 name f = Tac2env.define_primitive (pname name) begin fun bt arg -> match arg with +| [x; y; z] -> f bt x y z | _ -> assert false end (** Printing *) -let () = define1 "print" begin fun pp -> +let () = define1 "print" begin fun _ pp -> wrap_unit (fun () -> Feedback.msg_notice (Value.to_pp pp)) end -let () = define1 "message_of_int" begin fun n -> +let () = define1 "message_of_int" begin fun _ n -> let n = Value.to_int n in return (Value.of_pp (int n)) end -let () = define1 "message_of_string" begin fun s -> +let () = define1 "message_of_string" begin fun _ s -> let s = Value.to_string s in return (Value.of_pp (str (Bytes.to_string s))) end -let () = define1 "message_of_constr" begin fun c -> - pf_apply begin fun env sigma -> +let () = define1 "message_of_constr" begin fun bt c -> + pf_apply bt begin fun env sigma -> let c = Value.to_constr c in let pp = Printer.pr_econstr_env env sigma c in return (Value.of_pp pp) end end -let () = define1 "message_of_ident" begin fun c -> +let () = define1 "message_of_ident" begin fun _ c -> let c = Value.to_ident c in let pp = Id.print c in return (Value.of_pp pp) end -let () = define2 "message_concat" begin fun m1 m2 -> +let () = define2 "message_concat" begin fun _ m1 m2 -> let m1 = Value.to_pp m1 in let m2 = Value.to_pp m2 in return (Value.of_pp (Pp.app m1 m2)) @@ -194,45 +197,45 @@ end (** Array *) -let () = define2 "array_make" begin fun n x -> +let () = define2 "array_make" begin fun bt n x -> let n = Value.to_int n in - if n < 0 || n > Sys.max_array_length then throw err_outofbounds + if n < 0 || n > Sys.max_array_length then throw bt err_outofbounds else wrap (fun () -> ValBlk (0, Array.make n x)) end -let () = define1 "array_length" begin fun v -> +let () = define1 "array_length" begin fun _ v -> let v = to_block v in return (ValInt (Array.length v)) end -let () = define3 "array_set" begin fun v n x -> +let () = define3 "array_set" begin fun bt v n x -> let v = to_block v in let n = Value.to_int n in - if n < 0 || n >= Array.length v then throw err_outofbounds + if n < 0 || n >= Array.length v then throw bt err_outofbounds else wrap_unit (fun () -> v.(n) <- x) end -let () = define2 "array_get" begin fun v n -> +let () = define2 "array_get" begin fun bt v n -> let v = to_block v in let n = Value.to_int n in - if n < 0 || n >= Array.length v then throw err_outofbounds + if n < 0 || n >= Array.length v then throw bt err_outofbounds else wrap (fun () -> v.(n)) end (** Ident *) -let () = define2 "ident_equal" begin fun id1 id2 -> +let () = define2 "ident_equal" begin fun _ id1 id2 -> let id1 = Value.to_ident id1 in let id2 = Value.to_ident id2 in return (Value.of_bool (Id.equal id1 id2)) end -let () = define1 "ident_to_string" begin fun id -> +let () = define1 "ident_to_string" begin fun _ id -> let id = Value.to_ident id in return (Value.of_string (Id.to_string id)) end -let () = define1 "ident_of_string" begin fun s -> +let () = define1 "ident_of_string" begin fun _ s -> let s = Value.to_string s in let id = try Some (Id.of_string s) with _ -> None in return (Value.of_option Value.of_ident id) @@ -240,11 +243,11 @@ end (** Int *) -let () = define2 "int_equal" begin fun m n -> +let () = define2 "int_equal" begin fun _ m n -> return (Value.of_bool (Value.to_int m == Value.to_int n)) end -let binop n f = define2 n begin fun m n -> +let binop n f = define2 n begin fun _ m n -> return (Value.of_int (f (Value.to_int m) (Value.to_int n))) end @@ -253,42 +256,42 @@ let () = binop "int_add" (+) let () = binop "int_sub" (-) let () = binop "int_mul" ( * ) -let () = define1 "int_neg" begin fun m -> +let () = define1 "int_neg" begin fun _ m -> return (Value.of_int (~- (Value.to_int m))) end (** String *) -let () = define2 "string_make" begin fun n c -> +let () = define2 "string_make" begin fun bt n c -> let n = Value.to_int n in let c = Value.to_char c in - if n < 0 || n > Sys.max_string_length then throw err_outofbounds + if n < 0 || n > Sys.max_string_length then throw bt err_outofbounds else wrap (fun () -> Value.of_string (Bytes.make n c)) end -let () = define1 "string_length" begin fun s -> +let () = define1 "string_length" begin fun _ s -> return (Value.of_int (Bytes.length (Value.to_string s))) end -let () = define3 "string_set" begin fun s n c -> +let () = define3 "string_set" begin fun bt s n c -> let s = Value.to_string s in let n = Value.to_int n in let c = Value.to_char c in - if n < 0 || n >= Bytes.length s then throw err_outofbounds + if n < 0 || n >= Bytes.length s then throw bt err_outofbounds else wrap_unit (fun () -> Bytes.set s n c) end -let () = define2 "string_get" begin fun s n -> +let () = define2 "string_get" begin fun bt s n -> let s = Value.to_string s in let n = Value.to_int n in - if n < 0 || n >= Bytes.length s then throw err_outofbounds + if n < 0 || n >= Bytes.length s then throw bt err_outofbounds else wrap (fun () -> Value.of_char (Bytes.get s n)) end (** Terms *) (** constr -> constr *) -let () = define1 "constr_type" begin fun c -> +let () = define1 "constr_type" begin fun bt c -> let c = Value.to_constr c in let get_type env sigma = Proofview.V82.wrap_exceptions begin fun () -> @@ -296,11 +299,11 @@ let () = define1 "constr_type" begin fun c -> let t = Value.of_constr t in Proofview.Unsafe.tclEVARS sigma <*> Proofview.tclUNIT t end in - pf_apply get_type + pf_apply bt get_type end (** constr -> constr *) -let () = define2 "constr_equal" begin fun c1 c2 -> +let () = define2 "constr_equal" begin fun _ c1 c2 -> let c1 = Value.to_constr c1 in let c2 = Value.to_constr c2 in Proofview.tclEVARMAP >>= fun sigma -> @@ -308,7 +311,7 @@ let () = define2 "constr_equal" begin fun c1 c2 -> Proofview.tclUNIT (Value.of_bool b) end -let () = define1 "constr_kind" begin fun c -> +let () = define1 "constr_kind" begin fun _ c -> let open Constr in Proofview.tclEVARMAP >>= fun sigma -> let c = Value.to_constr c in @@ -403,7 +406,7 @@ let () = define1 "constr_kind" begin fun c -> end end -let () = define1 "constr_make" begin fun knd -> +let () = define1 "constr_make" begin fun _ knd -> let open Constr in let c = match knd with | ValBlk (0, [|n|]) -> @@ -483,9 +486,9 @@ let () = define1 "constr_make" begin fun knd -> return (Value.of_constr c) end -let () = define1 "constr_check" begin fun c -> +let () = define1 "constr_check" begin fun bt c -> let c = Value.to_constr c in - pf_apply begin fun env sigma -> + pf_apply bt begin fun env sigma -> try let (sigma, _) = Typing.type_of env sigma c in Proofview.Unsafe.tclEVARS sigma >>= fun () -> @@ -496,7 +499,7 @@ let () = define1 "constr_check" begin fun c -> end end -let () = define3 "constr_substnl" begin fun subst k c -> +let () = define3 "constr_substnl" begin fun _ subst k c -> let subst = Value.to_list Value.to_constr subst in let k = Value.to_int k in let c = Value.to_constr c in @@ -504,7 +507,7 @@ let () = define3 "constr_substnl" begin fun subst k c -> return (Value.of_constr ans) end -let () = define3 "constr_closenl" begin fun ids k c -> +let () = define3 "constr_closenl" begin fun _ ids k c -> let ids = Value.to_list Value.to_ident ids in let k = Value.to_int k in let c = Value.to_constr c in @@ -514,16 +517,16 @@ end (** Patterns *) -let () = define2 "pattern_matches" begin fun pat c -> +let () = define2 "pattern_matches" begin fun bt pat c -> let pat = Value.to_pattern pat in let c = Value.to_constr c in - pf_apply begin fun env sigma -> + pf_apply bt begin fun env sigma -> let ans = try Some (Constr_matching.matches env sigma pat c) with Constr_matching.PatternMatchingFailure -> None in begin match ans with - | None -> Proofview.tclZERO err_matchfailure + | None -> Proofview.tclZERO (err_matchfailure bt) | Some ans -> let ans = Id.Map.bindings ans in let of_pair (id, c) = Value.of_tuple [| Value.of_ident id; Value.of_constr c |] in @@ -532,34 +535,34 @@ let () = define2 "pattern_matches" begin fun pat c -> end end -let () = define2 "pattern_matches_subterm" begin fun pat c -> +let () = define2 "pattern_matches_subterm" begin fun bt pat c -> let pat = Value.to_pattern pat in let c = Value.to_constr c in let open Constr_matching in let rec of_ans s = match IStream.peek s with - | IStream.Nil -> Proofview.tclZERO err_matchfailure + | IStream.Nil -> Proofview.tclZERO (err_matchfailure bt) | IStream.Cons ({ m_sub = (_, sub); m_ctx }, s) -> let ans = Id.Map.bindings sub in let of_pair (id, c) = Value.of_tuple [| Value.of_ident id; Value.of_constr c |] in let ans = Value.of_tuple [| Value.of_constr m_ctx; Value.of_list of_pair ans |] in Proofview.tclOR (return ans) (fun _ -> of_ans s) in - pf_apply begin fun env sigma -> + pf_apply bt begin fun env sigma -> let ans = Constr_matching.match_appsubterm env sigma pat c in of_ans ans end end -let () = define2 "pattern_matches_vect" begin fun pat c -> +let () = define2 "pattern_matches_vect" begin fun bt pat c -> let pat = Value.to_pattern pat in let c = Value.to_constr c in - pf_apply begin fun env sigma -> + pf_apply bt begin fun env sigma -> let ans = try Some (Constr_matching.matches env sigma pat c) with Constr_matching.PatternMatchingFailure -> None in begin match ans with - | None -> Proofview.tclZERO err_matchfailure + | None -> Proofview.tclZERO (err_matchfailure bt) | Some ans -> let ans = Id.Map.bindings ans in let ans = Array.map_of_list snd ans in @@ -568,25 +571,25 @@ let () = define2 "pattern_matches_vect" begin fun pat c -> end end -let () = define2 "pattern_matches_subterm_vect" begin fun pat c -> +let () = define2 "pattern_matches_subterm_vect" begin fun bt pat c -> let pat = Value.to_pattern pat in let c = Value.to_constr c in let open Constr_matching in let rec of_ans s = match IStream.peek s with - | IStream.Nil -> Proofview.tclZERO err_matchfailure + | IStream.Nil -> Proofview.tclZERO (err_matchfailure bt) | IStream.Cons ({ m_sub = (_, sub); m_ctx }, s) -> let ans = Id.Map.bindings sub in let ans = Array.map_of_list snd ans in let ans = Value.of_tuple [| Value.of_constr m_ctx; Value.of_array Value.of_constr ans |] in Proofview.tclOR (return ans) (fun _ -> of_ans s) in - pf_apply begin fun env sigma -> + pf_apply bt begin fun env sigma -> let ans = Constr_matching.match_appsubterm env sigma pat c in of_ans ans end end -let () = define2 "pattern_instantiate" begin fun ctx c -> +let () = define2 "pattern_instantiate" begin fun _ ctx c -> let ctx = EConstr.Unsafe.to_constr (Value.to_constr ctx) in let c = EConstr.Unsafe.to_constr (Value.to_constr c) in let ans = Termops.subst_meta [Constr_matching.special_meta, c] ctx in @@ -595,46 +598,48 @@ end (** Error *) -let () = define1 "throw" begin fun e -> +let () = define1 "throw" begin fun bt e -> let (e, info) = Value.to_exn e in + let e = set_bt bt e in Proofview.tclLIFT (Proofview.NonLogical.raise ~info e) end (** Control *) (** exn -> 'a *) -let () = define1 "zero" begin fun e -> +let () = define1 "zero" begin fun bt e -> let (e, info) = Value.to_exn e in + let e = set_bt bt e in Proofview.tclZERO ~info e end (** (unit -> 'a) -> (exn -> 'a) -> 'a *) -let () = define2 "plus" begin fun x k -> - Proofview.tclOR (thaw x) (fun e -> interp_app k [Value.of_exn e]) +let () = define2 "plus" begin fun bt x k -> + Proofview.tclOR (thaw bt x) (fun e -> Tac2interp.interp_app bt k [Value.of_exn e]) end (** (unit -> 'a) -> 'a *) -let () = define1 "once" begin fun f -> - Proofview.tclONCE (thaw f) +let () = define1 "once" begin fun bt f -> + Proofview.tclONCE (thaw bt f) end (** (unit -> unit) list -> unit *) -let () = define1 "dispatch" begin fun l -> - let l = Value.to_list (fun f -> Proofview.tclIGNORE (thaw f)) l in +let () = define1 "dispatch" begin fun bt l -> + let l = Value.to_list (fun f -> Proofview.tclIGNORE (thaw bt f)) l in Proofview.tclDISPATCH l >>= fun () -> return v_unit end (** (unit -> unit) list -> (unit -> unit) -> (unit -> unit) list -> unit *) -let () = define3 "extend" begin fun lft tac rgt -> - let lft = Value.to_list (fun f -> Proofview.tclIGNORE (thaw f)) lft in - let tac = Proofview.tclIGNORE (thaw tac) in - let rgt = Value.to_list (fun f -> Proofview.tclIGNORE (thaw f)) rgt in +let () = define3 "extend" begin fun bt lft tac rgt -> + let lft = Value.to_list (fun f -> Proofview.tclIGNORE (thaw bt f)) lft in + let tac = Proofview.tclIGNORE (thaw bt tac) in + let rgt = Value.to_list (fun f -> Proofview.tclIGNORE (thaw bt f)) rgt in Proofview.tclEXTEND lft tac rgt >>= fun () -> return v_unit end (** (unit -> unit) -> unit *) -let () = define1 "enter" begin fun f -> - let f = Proofview.tclIGNORE (thaw f) in +let () = define1 "enter" begin fun bt f -> + let f = Proofview.tclIGNORE (thaw bt f) in Proofview.tclINDEPENDENT f >>= fun () -> return v_unit end @@ -643,8 +648,8 @@ let e_var = Id.of_string "e" let prm_apply_kont_h = pname "apply_kont" (** (unit -> 'a) -> ('a * ('exn -> 'a)) result *) -let () = define1 "case" begin fun f -> - Proofview.tclCASE (thaw f) >>= begin function +let () = define1 "case" begin fun bt f -> + Proofview.tclCASE (thaw bt f) >>= begin function | Proofview.Next (x, k) -> let k = { clos_ref = None; @@ -658,38 +663,40 @@ let () = define1 "case" begin fun f -> end (** 'a kont -> exn -> 'a *) -let () = define2 "apply_kont" begin fun k e -> - (Value.to_ext Value.val_kont k) (Value.to_exn e) +let () = define2 "apply_kont" begin fun bt k e -> + let (e, info) = Value.to_exn e in + let e = set_bt bt e in + (Value.to_ext Value.val_kont k) (e, info) end (** int -> int -> (unit -> 'a) -> 'a *) -let () = define3 "focus" begin fun i j tac -> +let () = define3 "focus" begin fun bt i j tac -> let i = Value.to_int i in let j = Value.to_int j in - Proofview.tclFOCUS i j (thaw tac) + Proofview.tclFOCUS i j (thaw bt tac) end (** unit -> unit *) -let () = define0 "shelve" begin +let () = define0 "shelve" begin fun _ -> Proofview.shelve >>= fun () -> return v_unit end (** unit -> unit *) -let () = define0 "shelve_unifiable" begin +let () = define0 "shelve_unifiable" begin fun _ -> Proofview.shelve_unifiable >>= fun () -> return v_unit end -let () = define1 "new_goal" begin fun ev -> +let () = define1 "new_goal" begin fun bt ev -> let ev = Evar.unsafe_of_int (Value.to_int ev) in Proofview.tclEVARMAP >>= fun sigma -> if Evd.mem sigma ev then Proofview.Unsafe.tclNEWGOALS [ev] <*> Proofview.tclUNIT v_unit - else throw err_notfound + else throw bt err_notfound end (** unit -> constr *) -let () = define0 "goal" begin - assert_focussed >>= fun () -> +let () = define0 "goal" begin fun bt -> + assert_focussed bt >>= fun () -> Proofview.Goal.enter_one begin fun gl -> let concl = Tacmach.New.pf_nf_concl gl in return (Value.of_constr concl) @@ -697,9 +704,9 @@ let () = define0 "goal" begin end (** ident -> constr *) -let () = define1 "hyp" begin fun id -> +let () = define1 "hyp" begin fun bt id -> let id = Value.to_ident id in - pf_apply begin fun env _ -> + pf_apply bt begin fun env _ -> let mem = try ignore (Environ.lookup_named id env); true with Not_found -> false in if mem then return (Value.of_constr (EConstr.mkVar id)) else Tacticals.New.tclZEROMSG @@ -707,8 +714,8 @@ let () = define1 "hyp" begin fun id -> end end -let () = define0 "hyps" begin - pf_apply begin fun env _ -> +let () = define0 "hyps" begin fun bt -> + pf_apply bt begin fun env _ -> let open Context.Named.Declaration in let hyps = List.rev (Environ.named_context env) in let map = function @@ -725,56 +732,56 @@ let () = define0 "hyps" begin end (** (unit -> constr) -> unit *) -let () = define1 "refine" begin fun c -> - let c = thaw c >>= fun c -> Proofview.tclUNIT ((), Value.to_constr c) in +let () = define1 "refine" begin fun bt c -> + let c = thaw bt c >>= fun c -> Proofview.tclUNIT ((), Value.to_constr c) in Proofview.Goal.nf_enter begin fun gl -> Refine.generic_refine ~typecheck:true c gl end >>= fun () -> return v_unit end -let () = define2 "with_holes" begin fun x f -> +let () = define2 "with_holes" begin fun bt x f -> Proofview.tclEVARMAP >>= fun sigma0 -> - thaw x >>= fun ans -> + thaw bt x >>= fun ans -> Proofview.tclEVARMAP >>= fun sigma -> Proofview.Unsafe.tclEVARS sigma0 >>= fun () -> - Tacticals.New.tclWITHHOLES false (interp_app f [ans]) sigma + Tacticals.New.tclWITHHOLES false (Tac2interp.interp_app bt f [ans]) sigma end -let () = define1 "progress" begin fun f -> - Proofview.tclPROGRESS (thaw f) +let () = define1 "progress" begin fun bt f -> + Proofview.tclPROGRESS (thaw bt f) end -let () = define2 "abstract" begin fun id f -> +let () = define2 "abstract" begin fun bt id f -> let id = Value.to_option Value.to_ident id in - Tactics.tclABSTRACT id (Proofview.tclIGNORE (thaw f)) >>= fun () -> + Tactics.tclABSTRACT id (Proofview.tclIGNORE (thaw bt f)) >>= fun () -> return v_unit end -let () = define2 "time" begin fun s f -> +let () = define2 "time" begin fun bt s f -> let s = Value.to_option Value.to_string s in - Proofview.tclTIME s (thaw f) + Proofview.tclTIME s (thaw bt f) end -let () = define0 "check_interrupt" begin +let () = define0 "check_interrupt" begin fun bt -> Proofview.tclCHECKINTERRUPT >>= fun () -> return v_unit end (** Fresh *) -let () = define2 "fresh_free_union" begin fun set1 set2 -> +let () = define2 "fresh_free_union" begin fun _ set1 set2 -> let set1 = Value.to_ext Value.val_free set1 in let set2 = Value.to_ext Value.val_free set2 in let ans = Id.Set.union set1 set2 in return (Value.of_ext Value.val_free ans) end -let () = define1 "fresh_free_of_ids" begin fun ids -> +let () = define1 "fresh_free_of_ids" begin fun _ ids -> let ids = Value.to_list Value.to_ident ids in let free = List.fold_right Id.Set.add ids Id.Set.empty in return (Value.of_ext Value.val_free free) end -let () = define1 "fresh_free_of_constr" begin fun c -> +let () = define1 "fresh_free_of_constr" begin fun _ c -> let c = Value.to_constr c in Proofview.tclEVARMAP >>= fun sigma -> let rec fold accu c = match EConstr.kind sigma c with @@ -785,7 +792,7 @@ let () = define1 "fresh_free_of_constr" begin fun c -> return (Value.of_ext Value.val_free ans) end -let () = define2 "fresh_fresh" begin fun avoid id -> +let () = define2 "fresh_fresh" begin fun _ avoid id -> let avoid = Value.to_ext Value.val_free avoid in let id = Value.to_ident id in let nid = Namegen.next_ident_away_from id (fun id -> Id.Set.mem id avoid) in @@ -828,9 +835,10 @@ let intern_constr self ist c = let interp_constr flags ist c = let open Pretyping in - pf_apply begin fun env sigma -> + let bt = ist.env_bkt in + let ist = to_lvar ist in + pf_apply bt begin fun env sigma -> Proofview.V82.wrap_exceptions begin fun () -> - let ist = to_lvar ist in let (sigma, c) = understand_ltac flags env sigma ist WithoutTypeConstraint c in let c = ValExt (Value.val_constr, c) in Proofview.Unsafe.tclEVARS sigma >>= fun () -> @@ -944,7 +952,7 @@ let () = let () = let interp ist env sigma concl tac = let ist = Tac2interp.get_env ist in - let tac = Proofview.tclIGNORE (interp ist tac) in + let tac = Proofview.tclIGNORE (Tac2interp.interp ist tac) in let c, sigma = Pfedit.refine_by_tactic env sigma concl tac in (EConstr.of_constr c, sigma) in @@ -965,7 +973,9 @@ let () = (** FUCK YOU API *) let idtac = (Obj.magic idtac : Geninterp.Val.t) in let interp ist tac = - Tac2interp.interp Tac2interp.empty_environment tac >>= fun _ -> + let ist = Tac2interp.get_env ist.Geninterp.lfun in + let ist = { ist with env_ist = Id.Map.empty } in + Tac2interp.interp ist tac >>= fun _ -> Ftactic.return idtac in Geninterp.register_interp0 wit_ltac2 interp diff --git a/src/tac2core.mli b/src/tac2core.mli index 9fae65bb3e..b5800a7172 100644 --- a/src/tac2core.mli +++ b/src/tac2core.mli @@ -27,4 +27,4 @@ val c_false : ltac_constructor end -val pf_apply : (Environ.env -> Evd.evar_map -> 'a Proofview.tactic) -> 'a Proofview.tactic +val pf_apply : backtrace -> (Environ.env -> Evd.evar_map -> 'a Proofview.tactic) -> 'a Proofview.tactic diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 0754108505..8d515577ec 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -743,6 +743,41 @@ let register_struct ?local str = match str with | StrSyn (tok, lev, e) -> register_notation ?local tok lev e | StrMut (qid, e) -> register_redefinition ?local qid e +(** Toplevel exception *) + +let print_ltac2_backtrace = ref false + +let _ = Goptions.declare_bool_option { + Goptions.optdepr = false; + Goptions.optname = "print Ltac2 backtrace"; + Goptions.optkey = ["Ltac2"; "Backtrace"]; + Goptions.optread = (fun () -> !print_ltac2_backtrace); + Goptions.optwrite = (fun b -> print_ltac2_backtrace := b); +} + +let pr_frame = function +| FrLtac None -> str "" +| FrLtac (Some kn) -> + Libnames.pr_qualid (Tac2env.shortest_qualid_of_ltac (TacConstant kn)) +| FrPrim ml -> + str "<" ++ str ml.mltac_plugin ++ str ":" ++ str ml.mltac_tactic ++ str ">" +| FrExtn (tag, arg) -> + let obj = Tac2env.interp_ml_object tag in + obj.Tac2env.ml_print (Global.env ()) arg + +let () = register_handler begin function +| Tac2interp.LtacError (kn, _, bt) -> + let c = Tac2print.pr_constructor kn in (** FIXME *) + let bt = + if !print_ltac2_backtrace then + fnl () ++ str "Backtrace:" ++ fnl () ++ prlist_with_sep fnl pr_frame bt + else + mt () + in + hov 0 (str "Uncaught Ltac2 exception:" ++ spc () ++ hov 0 c) ++ bt +| _ -> raise Unhandled +end + (** Printing *) let print_ltac ref = diff --git a/src/tac2expr.mli b/src/tac2expr.mli index 470323e7c7..36c3fbbe59 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -205,6 +205,16 @@ and closure = { (** Global constant from which the closure originates *) } -type ml_tactic = valexpr list -> valexpr Proofview.tactic +type frame = +| FrLtac of ltac_constant option +| FrPrim of ml_tactic_name +| FrExtn : ('a, 'b) Tac2dyn.Arg.tag * 'b -> frame -type environment = valexpr Id.Map.t +type backtrace = frame list + +type ml_tactic = backtrace -> valexpr list -> valexpr Proofview.tactic + +type environment = { + env_ist : valexpr Id.Map.t; + env_bkt : backtrace; +} diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index dd20de5ef5..4ed0096787 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -11,7 +11,6 @@ open Globnames open Genarg open Tac2dyn open Tac2expr -open Tac2interp (** Dynamic tags *) @@ -98,7 +97,7 @@ let internal_err = (** FIXME: handle backtrace in Ltac2 exceptions *) let of_exn c = match fst c with -| LtacError (kn, c) -> ValOpn (kn, c) +| Tac2interp.LtacError (kn, c, _) -> ValOpn (kn, c) | _ -> ValOpn (internal_err, [|of_ext val_exn c|]) let to_exn c = match c with @@ -106,7 +105,7 @@ let to_exn c = match c with if Names.KerName.equal kn internal_err then to_ext val_exn c.(0) else - (LtacError (kn, c), Exninfo.null) + (Tac2interp.LtacError (kn, c, []), Exninfo.null) | _ -> assert false let of_option f = function diff --git a/src/tac2interp.ml b/src/tac2interp.ml index 7bcfad1be1..b58ce6b851 100644 --- a/src/tac2interp.ml +++ b/src/tac2interp.ml @@ -14,25 +14,19 @@ open Names open Proofview.Notations open Tac2expr -exception LtacError of KerName.t * valexpr array +exception LtacError of KerName.t * valexpr array * backtrace -let () = register_handler begin function -| LtacError (kn, _) -> - let c = Tac2print.pr_constructor kn in - hov 0 (str "Uncaught Ltac2 exception:" ++ spc () ++ hov 0 c) -| _ -> raise Unhandled -end - -type environment = valexpr Id.Map.t - -let empty_environment = Id.Map.empty +let empty_environment = { + env_ist = Id.Map.empty; + env_bkt = []; +} let push_name ist id v = match id with | Anonymous -> ist -| Name id -> Id.Map.add id v ist +| Name id -> { ist with env_ist = Id.Map.add id v ist.env_ist } let get_var ist id = - try Id.Map.find id ist with Not_found -> + try Id.Map.find id ist.env_ist with Not_found -> anomaly (str "Unbound variable " ++ Id.print id) let get_ref ist kn = @@ -41,18 +35,18 @@ let get_ref ist kn = let return = Proofview.tclUNIT -let rec interp ist = function +let rec interp (ist : environment) = function | GTacAtm (AtmInt n) -> return (ValInt n) | GTacAtm (AtmStr s) -> return (ValStr (Bytes.of_string s)) | GTacVar id -> return (get_var ist id) | GTacRef qid -> return (get_ref ist qid) | GTacFun (ids, e) -> - let cls = { clos_ref = None; clos_env = ist; clos_var = ids; clos_exp = e } in + let cls = { clos_ref = None; clos_env = ist.env_ist; clos_var = ids; clos_exp = e } in return (ValCls cls) | GTacApp (f, args) -> interp ist f >>= fun f -> Proofview.Monad.List.map (fun e -> interp ist e) args >>= fun args -> - interp_app f args + interp_app ist.env_bkt f args | GTacLet (false, el, e) -> let fold accu (na, e) = interp ist e >>= fun e -> @@ -63,18 +57,18 @@ let rec interp ist = function | GTacLet (true, el, e) -> let map (na, e) = match e with | GTacFun (ids, e) -> - let cls = { clos_ref = None; clos_env = ist; clos_var = ids; clos_exp = e } in + let cls = { clos_ref = None; clos_env = ist.env_ist; clos_var = ids; clos_exp = e } in na, cls | _ -> anomaly (str "Ill-formed recursive function") in let fixs = List.map map el in let fold accu (na, cls) = match na with | Anonymous -> accu - | Name id -> Id.Map.add id (ValCls cls) accu + | Name id -> { ist with env_ist = Id.Map.add id (ValCls cls) accu.env_ist } in let ist = List.fold_left fold ist fixs in (** Hack to make a cycle imperatively in the environment *) - let iter (_, e) = e.clos_env <- ist in + let iter (_, e) = e.clos_env <- ist.env_ist in let () = List.iter iter fixs in interp ist e | GTacCst (_, n, []) -> return (ValInt n) @@ -96,22 +90,23 @@ let rec interp ist = function return (ValOpn (kn, Array.of_list el)) | GTacPrm (ml, el) -> Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> - Tac2env.interp_primitive ml el + Tac2env.interp_primitive ml (FrPrim ml :: ist.env_bkt) el | GTacExt (tag, e) -> let tpe = Tac2env.interp_ml_object tag in + let ist = { ist with env_bkt = FrExtn (tag, e) :: ist.env_bkt } in tpe.Tac2env.ml_interp ist e -and interp_app f args = match f with +and interp_app bt f args = match f with | ValCls { clos_env = ist; clos_var = ids; clos_exp = e; clos_ref = kn } -> let rec push ist ids args = match ids, args with | [], [] -> interp ist e - | [], _ :: _ -> interp ist e >>= fun f -> interp_app f args + | [], _ :: _ -> interp ist e >>= fun f -> interp_app bt f args | _ :: _, [] -> - let cls = { clos_ref = kn; clos_env = ist; clos_var = ids; clos_exp = e } in + let cls = { clos_ref = kn; clos_env = ist.env_ist; clos_var = ids; clos_exp = e } in return (ValCls cls) | id :: ids, arg :: args -> push (push_name ist id arg) ids args in - push ist ids args + push { env_ist = ist; env_bkt = FrLtac kn :: bt } ids args | ValExt _ | ValInt _ | ValBlk _ | ValStr _ | ValOpn _ -> anomaly (str "Unexpected value shape") diff --git a/src/tac2interp.mli b/src/tac2interp.mli index 1ac26b48db..ea7db33b60 100644 --- a/src/tac2interp.mli +++ b/src/tac2interp.mli @@ -9,13 +9,11 @@ open Names open Tac2expr -type environment = valexpr Id.Map.t - val empty_environment : environment val interp : environment -> glb_tacexpr -> valexpr Proofview.tactic -val interp_app : valexpr -> valexpr list -> valexpr Proofview.tactic +val interp_app : backtrace -> valexpr -> valexpr list -> valexpr Proofview.tactic (** {5 Cross-boundary encodings} *) @@ -24,5 +22,5 @@ val set_env : environment -> Glob_term.unbound_ltac_var_map -> Glob_term.unbound (** {5 Exceptions} *) -exception LtacError of KerName.t * valexpr array +exception LtacError of KerName.t * valexpr array * backtrace (** Ltac2-defined exceptions seen from OCaml side *) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 2a57fdc705..6dcb3f15fb 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -18,7 +18,7 @@ module Value = Tac2ffi let return x = Proofview.tclUNIT x let v_unit = Value.of_unit () -let thaw f = Tac2interp.interp_app f [v_unit] +let thaw bt f = Tac2interp.interp_app bt f [v_unit] let to_pair f g = function | ValBlk (0, [| x; y |]) -> (f x, g y) @@ -126,7 +126,8 @@ and to_intro_patterns il = let to_destruction_arg = function | ValBlk (0, [| c |]) -> - let c = thaw c >>= fun c -> return (to_constr_with_bindings c) in + (** FIXME: lost backtrace *) + let c = thaw [] c >>= fun c -> return (to_constr_with_bindings c) in ElimOnConstr c | ValBlk (1, [| id |]) -> ElimOnIdent (Loc.tag (Value.to_ident id)) | ValBlk (2, [| n |]) -> ElimOnAnonHyp (Value.to_int n) @@ -153,7 +154,8 @@ let to_rewriting = function | ValBlk (0, [| orient; repeat; c |]) -> let orient = Value.to_option Value.to_bool orient in let repeat = to_multi repeat in - let c = thaw c >>= fun c -> return (to_constr_with_bindings c) in + (** FIXME: lost backtrace *) + let c = thaw [] c >>= fun c -> return (to_constr_with_bindings c) in (orient, repeat, c) | _ -> assert false @@ -164,52 +166,52 @@ let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } let lift tac = tac <*> return v_unit let define_prim0 name tac = - let tac = function + let tac bt arg = match arg with | [_] -> lift tac | _ -> assert false in Tac2env.define_primitive (pname name) tac let define_prim1 name tac = - let tac = function - | [x] -> lift (tac x) + let tac bt arg = match arg with + | [x] -> lift (tac bt x) | _ -> assert false in Tac2env.define_primitive (pname name) tac let define_prim2 name tac = - let tac = function - | [x; y] -> lift (tac x y) + let tac bt arg = match arg with + | [x; y] -> lift (tac bt x y) | _ -> assert false in Tac2env.define_primitive (pname name) tac let define_prim3 name tac = - let tac = function - | [x; y; z] -> lift (tac x y z) + let tac bt arg = match arg with + | [x; y; z] -> lift (tac bt x y z) | _ -> assert false in Tac2env.define_primitive (pname name) tac let define_prim4 name tac = - let tac = function - | [x; y; z; u] -> lift (tac x y z u) + let tac bt arg = match arg with + | [x; y; z; u] -> lift (tac bt x y z u) | _ -> assert false in Tac2env.define_primitive (pname name) tac (** Tactics from Tacexpr *) -let () = define_prim2 "tac_intros" begin fun ev ipat -> +let () = define_prim2 "tac_intros" begin fun _ ev ipat -> let ev = Value.to_bool ev in let ipat = to_intro_patterns ipat in Tactics.intros_patterns ev ipat end -let () = define_prim4 "tac_apply" begin fun adv ev cb ipat -> +let () = define_prim4 "tac_apply" begin fun bt adv ev cb ipat -> let adv = Value.to_bool adv in let ev = Value.to_bool ev in - let map_cb c = thaw c >>= fun c -> return (to_constr_with_bindings c) in + let map_cb c = thaw bt c >>= fun c -> return (to_constr_with_bindings c) in let cb = Value.to_list map_cb cb in let map p = Value.to_option (fun p -> Loc.tag (to_intro_pattern p)) p in let map_ipat p = to_pair Value.to_ident map p in @@ -217,20 +219,20 @@ let () = define_prim4 "tac_apply" begin fun adv ev cb ipat -> Tac2tactics.apply adv ev cb ipat end -let () = define_prim3 "tac_elim" begin fun ev c copt -> +let () = define_prim3 "tac_elim" begin fun _ ev c copt -> let ev = Value.to_bool ev in let c = to_constr_with_bindings c in let copt = Value.to_option to_constr_with_bindings copt in Tactics.elim ev None c copt end -let () = define_prim2 "tac_case" begin fun ev c -> +let () = define_prim2 "tac_case" begin fun _ ev c -> let ev = Value.to_bool ev in let c = to_constr_with_bindings c in Tactics.general_case_analysis ev None c end -let () = define_prim1 "tac_generalize" begin fun cl -> +let () = define_prim1 "tac_generalize" begin fun _ cl -> let cast = function | ValBlk (0, [| c; occs; na |]) -> ((to_occurrences Value.to_int occs, Value.to_constr c), to_name na) @@ -240,113 +242,113 @@ let () = define_prim1 "tac_generalize" begin fun cl -> Tactics.new_generalize_gen cl end -let () = define_prim3 "tac_assert" begin fun c tac ipat -> +let () = define_prim3 "tac_assert" begin fun bt c tac ipat -> let c = Value.to_constr c in - let of_tac t = Proofview.tclIGNORE (thaw t) in + let of_tac t = Proofview.tclIGNORE (thaw bt t) in let tac = Value.to_option (fun t -> Value.to_option of_tac t) tac in let ipat = Value.to_option (fun ipat -> Loc.tag (to_intro_pattern ipat)) ipat in Tactics.forward true tac ipat c end -let () = define_prim3 "tac_enough" begin fun c tac ipat -> +let () = define_prim3 "tac_enough" begin fun bt c tac ipat -> let c = Value.to_constr c in - let of_tac t = Proofview.tclIGNORE (thaw t) in + let of_tac t = Proofview.tclIGNORE (thaw bt t) in let tac = Value.to_option (fun t -> Value.to_option of_tac t) tac in let ipat = Value.to_option (fun ipat -> Loc.tag (to_intro_pattern ipat)) ipat in Tactics.forward false tac ipat c end -let () = define_prim2 "tac_pose" begin fun idopt c -> +let () = define_prim2 "tac_pose" begin fun _ idopt c -> let na = to_name idopt in let c = Value.to_constr c in Tactics.letin_tac None na c None Locusops.nowhere end -let () = define_prim4 "tac_set" begin fun ev idopt c cl -> +let () = define_prim4 "tac_set" begin fun bt ev idopt c cl -> let ev = Value.to_bool ev in let na = to_name idopt in let cl = to_clause cl in Proofview.tclEVARMAP >>= fun sigma -> - thaw c >>= fun c -> + thaw bt c >>= fun c -> let c = Value.to_constr c in Tactics.letin_pat_tac ev None na (sigma, c) cl end -let () = define_prim3 "tac_destruct" begin fun ev ic using -> +let () = define_prim3 "tac_destruct" begin fun _ ev ic using -> let ev = Value.to_bool ev in let ic = Value.to_list to_induction_clause ic in let using = Value.to_option to_constr_with_bindings using in Tac2tactics.induction_destruct false ev ic using end -let () = define_prim3 "tac_induction" begin fun ev ic using -> +let () = define_prim3 "tac_induction" begin fun _ ev ic using -> let ev = Value.to_bool ev in let ic = Value.to_list to_induction_clause ic in let using = Value.to_option to_constr_with_bindings using in Tac2tactics.induction_destruct true ev ic using end -let () = define_prim1 "tac_red" begin fun cl -> +let () = define_prim1 "tac_red" begin fun _ cl -> let cl = to_clause cl in Tactics.reduce (Red false) cl end -let () = define_prim1 "tac_hnf" begin fun cl -> +let () = define_prim1 "tac_hnf" begin fun _ cl -> let cl = to_clause cl in Tactics.reduce Hnf cl end -let () = define_prim3 "tac_simpl" begin fun flags where cl -> +let () = define_prim3 "tac_simpl" begin fun _ flags where cl -> let flags = to_red_flag flags in let where = Value.to_option to_pattern_with_occs where in let cl = to_clause cl in Tac2tactics.simpl flags where cl end -let () = define_prim2 "tac_cbv" begin fun flags cl -> +let () = define_prim2 "tac_cbv" begin fun _ flags cl -> let flags = to_red_flag flags in let cl = to_clause cl in Tac2tactics.cbv flags cl end -let () = define_prim2 "tac_cbn" begin fun flags cl -> +let () = define_prim2 "tac_cbn" begin fun _ flags cl -> let flags = to_red_flag flags in let cl = to_clause cl in Tac2tactics.cbn flags cl end -let () = define_prim2 "tac_lazy" begin fun flags cl -> +let () = define_prim2 "tac_lazy" begin fun _ flags cl -> let flags = to_red_flag flags in let cl = to_clause cl in Tac2tactics.lazy_ flags cl end -let () = define_prim2 "tac_unfold" begin fun refs cl -> +let () = define_prim2 "tac_unfold" begin fun _ refs cl -> let map v = to_pair Value.to_reference (fun occ -> to_occurrences to_int_or_var occ) v in let refs = Value.to_list map refs in let cl = to_clause cl in Tac2tactics.unfold refs cl end -let () = define_prim2 "tac_fold" begin fun args cl -> +let () = define_prim2 "tac_fold" begin fun _ args cl -> let args = Value.to_list Value.to_constr args in let cl = to_clause cl in Tactics.reduce (Fold args) cl end -let () = define_prim2 "tac_pattern" begin fun where cl -> +let () = define_prim2 "tac_pattern" begin fun _ where cl -> let where = Value.to_list to_constr_with_occs where in let cl = to_clause cl in Tactics.reduce (Pattern where) cl end -let () = define_prim2 "tac_vm" begin fun where cl -> +let () = define_prim2 "tac_vm" begin fun _ where cl -> let where = Value.to_option to_pattern_with_occs where in let cl = to_clause cl in Tac2tactics.vm where cl end -let () = define_prim2 "tac_native" begin fun where cl -> +let () = define_prim2 "tac_native" begin fun _ where cl -> let where = Value.to_option to_pattern_with_occs where in let cl = to_clause cl in Tac2tactics.native where cl @@ -355,97 +357,97 @@ end (** Reduction functions *) let define_red1 name tac = - let tac = function - | [x] -> tac x >>= fun c -> Proofview.tclUNIT (Value.of_constr c) + let tac bt arg = match arg with + | [x] -> tac bt x >>= fun c -> Proofview.tclUNIT (Value.of_constr c) | _ -> assert false in Tac2env.define_primitive (pname name) tac let define_red2 name tac = - let tac = function - | [x; y] -> tac x y >>= fun c -> Proofview.tclUNIT (Value.of_constr c) + let tac bt arg = match arg with + | [x; y] -> tac bt x y >>= fun c -> Proofview.tclUNIT (Value.of_constr c) | _ -> assert false in Tac2env.define_primitive (pname name) tac let define_red3 name tac = - let tac = function - | [x; y; z] -> tac x y z >>= fun c -> Proofview.tclUNIT (Value.of_constr c) + let tac bt arg = match arg with + | [x; y; z] -> tac bt x y z >>= fun c -> Proofview.tclUNIT (Value.of_constr c) | _ -> assert false in Tac2env.define_primitive (pname name) tac -let () = define_red1 "eval_red" begin fun c -> +let () = define_red1 "eval_red" begin fun bt c -> let c = Value.to_constr c in - Tac2tactics.eval_red c + Tac2tactics.eval_red bt c end -let () = define_red1 "eval_hnf" begin fun c -> +let () = define_red1 "eval_hnf" begin fun bt c -> let c = Value.to_constr c in - Tac2tactics.eval_hnf c + Tac2tactics.eval_hnf bt c end -let () = define_red3 "eval_simpl" begin fun flags where c -> +let () = define_red3 "eval_simpl" begin fun bt flags where c -> let flags = to_red_flag flags in let where = Value.to_option to_pattern_with_occs where in let c = Value.to_constr c in - Tac2tactics.eval_simpl flags where c + Tac2tactics.eval_simpl bt flags where c end -let () = define_red2 "eval_cbv" begin fun flags c -> +let () = define_red2 "eval_cbv" begin fun bt flags c -> let flags = to_red_flag flags in let c = Value.to_constr c in - Tac2tactics.eval_cbv flags c + Tac2tactics.eval_cbv bt flags c end -let () = define_red2 "eval_cbn" begin fun flags c -> +let () = define_red2 "eval_cbn" begin fun bt flags c -> let flags = to_red_flag flags in let c = Value.to_constr c in - Tac2tactics.eval_cbn flags c + Tac2tactics.eval_cbn bt flags c end -let () = define_red2 "eval_lazy" begin fun flags c -> +let () = define_red2 "eval_lazy" begin fun bt flags c -> let flags = to_red_flag flags in let c = Value.to_constr c in - Tac2tactics.eval_lazy flags c + Tac2tactics.eval_lazy bt flags c end -let () = define_red2 "eval_unfold" begin fun refs c -> +let () = define_red2 "eval_unfold" begin fun bt refs c -> let map v = to_pair Value.to_reference (fun occ -> to_occurrences to_int_or_var occ) v in let refs = Value.to_list map refs in let c = Value.to_constr c in - Tac2tactics.eval_unfold refs c + Tac2tactics.eval_unfold bt refs c end -let () = define_red2 "eval_fold" begin fun args c -> +let () = define_red2 "eval_fold" begin fun bt args c -> let args = Value.to_list Value.to_constr args in let c = Value.to_constr c in - Tac2tactics.eval_fold args c + Tac2tactics.eval_fold bt args c end -let () = define_red2 "eval_pattern" begin fun where c -> +let () = define_red2 "eval_pattern" begin fun bt where c -> let where = Value.to_list (fun p -> to_pair Value.to_constr (fun occ -> to_occurrences to_int_or_var occ) p) where in let c = Value.to_constr c in - Tac2tactics.eval_pattern where c + Tac2tactics.eval_pattern bt where c end -let () = define_red2 "eval_vm" begin fun where c -> +let () = define_red2 "eval_vm" begin fun bt where c -> let where = Value.to_option to_pattern_with_occs where in let c = Value.to_constr c in - Tac2tactics.eval_vm where c + Tac2tactics.eval_vm bt where c end -let () = define_red2 "eval_native" begin fun where c -> +let () = define_red2 "eval_native" begin fun bt where c -> let where = Value.to_option to_pattern_with_occs where in let c = Value.to_constr c in - Tac2tactics.eval_native where c + Tac2tactics.eval_native bt where c end -let () = define_prim4 "tac_rewrite" begin fun ev rw cl by -> +let () = define_prim4 "tac_rewrite" begin fun bt ev rw cl by -> let ev = Value.to_bool ev in let rw = Value.to_list to_rewriting rw in let cl = to_clause cl in - let to_tac t = Proofview.tclIGNORE (thaw t) in + let to_tac t = Proofview.tclIGNORE (thaw bt t) in let by = Value.to_option to_tac by in Tac2tactics.rewrite ev rw cl by end @@ -464,69 +466,69 @@ END let () = define_prim0 "tac_assumption" Tactics.assumption -let () = define_prim1 "tac_transitivity" begin fun c -> +let () = define_prim1 "tac_transitivity" begin fun _ c -> let c = Value.to_constr c in Tactics.intros_transitivity (Some c) end let () = define_prim0 "tac_etransitivity" (Tactics.intros_transitivity None) -let () = define_prim1 "tac_cut" begin fun c -> +let () = define_prim1 "tac_cut" begin fun _ c -> let c = Value.to_constr c in Tactics.cut c end -let () = define_prim2 "tac_left" begin fun ev bnd -> +let () = define_prim2 "tac_left" begin fun _ ev bnd -> let ev = Value.to_bool ev in let bnd = to_bindings bnd in Tactics.left_with_bindings ev bnd end -let () = define_prim2 "tac_right" begin fun ev bnd -> +let () = define_prim2 "tac_right" begin fun _ ev bnd -> let ev = Value.to_bool ev in let bnd = to_bindings bnd in Tactics.right_with_bindings ev bnd end -let () = define_prim1 "tac_introsuntil" begin fun h -> +let () = define_prim1 "tac_introsuntil" begin fun _ h -> Tactics.intros_until (to_qhyp h) end -let () = define_prim1 "tac_exactnocheck" begin fun c -> +let () = define_prim1 "tac_exactnocheck" begin fun _ c -> Tactics.exact_no_check (Value.to_constr c) end -let () = define_prim1 "tac_vmcastnocheck" begin fun c -> +let () = define_prim1 "tac_vmcastnocheck" begin fun _ c -> Tactics.vm_cast_no_check (Value.to_constr c) end -let () = define_prim1 "tac_nativecastnocheck" begin fun c -> +let () = define_prim1 "tac_nativecastnocheck" begin fun _ c -> Tactics.native_cast_no_check (Value.to_constr c) end -let () = define_prim1 "tac_constructor" begin fun ev -> +let () = define_prim1 "tac_constructor" begin fun _ ev -> let ev = Value.to_bool ev in Tactics.any_constructor ev None end -let () = define_prim3 "tac_constructorn" begin fun ev n bnd -> +let () = define_prim3 "tac_constructorn" begin fun _ ev n bnd -> let ev = Value.to_bool ev in let n = Value.to_int n in let bnd = to_bindings bnd in Tactics.constructor_tac ev None n bnd end -let () = define_prim1 "tac_symmetry" begin fun cl -> +let () = define_prim1 "tac_symmetry" begin fun _ cl -> let cl = to_clause cl in Tactics.intros_symmetry cl end -let () = define_prim2 "tac_split" begin fun ev bnd -> +let () = define_prim2 "tac_split" begin fun _ ev bnd -> let ev = Value.to_bool ev in let bnd = to_bindings bnd in Tactics.split_with_bindings ev [bnd] end -let () = define_prim1 "tac_rename" begin fun ids -> +let () = define_prim1 "tac_rename" begin fun _ ids -> let map c = match Value.to_tuple c with | [|x; y|] -> (Value.to_ident x, Value.to_ident y) | _ -> assert false @@ -535,72 +537,72 @@ let () = define_prim1 "tac_rename" begin fun ids -> Tactics.rename_hyp ids end -let () = define_prim1 "tac_revert" begin fun ids -> +let () = define_prim1 "tac_revert" begin fun _ ids -> let ids = Value.to_list Value.to_ident ids in Tactics.revert ids end let () = define_prim0 "tac_admit" Proofview.give_up -let () = define_prim2 "tac_fix" begin fun idopt n -> +let () = define_prim2 "tac_fix" begin fun _ idopt n -> let idopt = Value.to_option Value.to_ident idopt in let n = Value.to_int n in Tactics.fix idopt n end -let () = define_prim1 "tac_cofix" begin fun idopt -> +let () = define_prim1 "tac_cofix" begin fun _ idopt -> let idopt = Value.to_option Value.to_ident idopt in Tactics.cofix idopt end -let () = define_prim1 "tac_clear" begin fun ids -> +let () = define_prim1 "tac_clear" begin fun _ ids -> let ids = Value.to_list Value.to_ident ids in Tactics.clear ids end -let () = define_prim1 "tac_keep" begin fun ids -> +let () = define_prim1 "tac_keep" begin fun _ ids -> let ids = Value.to_list Value.to_ident ids in Tactics.keep ids end -let () = define_prim1 "tac_clearbody" begin fun ids -> +let () = define_prim1 "tac_clearbody" begin fun _ ids -> let ids = Value.to_list Value.to_ident ids in Tactics.clear_body ids end (** Tactics from extratactics *) -let () = define_prim2 "tac_discriminate" begin fun ev arg -> +let () = define_prim2 "tac_discriminate" begin fun _ ev arg -> let ev = Value.to_bool ev in let arg = Value.to_option (fun arg -> None, to_destruction_arg arg) arg in Tac2tactics.discriminate ev arg end -let () = define_prim3 "tac_injection" begin fun ev ipat arg -> +let () = define_prim3 "tac_injection" begin fun _ ev ipat arg -> let ev = Value.to_bool ev in let ipat = Value.to_option to_intro_patterns ipat in let arg = Value.to_option (fun arg -> None, to_destruction_arg arg) arg in Tac2tactics.injection ev ipat arg end -let () = define_prim1 "tac_absurd" begin fun c -> +let () = define_prim1 "tac_absurd" begin fun _ c -> Contradiction.absurd (Value.to_constr c) end -let () = define_prim1 "tac_contradiction" begin fun c -> +let () = define_prim1 "tac_contradiction" begin fun _ c -> let c = Value.to_option to_constr_with_bindings c in Contradiction.contradiction c end -let () = define_prim4 "tac_autorewrite" begin fun all by ids cl -> +let () = define_prim4 "tac_autorewrite" begin fun bt all by ids cl -> let all = Value.to_bool all in - let by = Value.to_option (fun tac -> Proofview.tclIGNORE (thaw tac)) by in + let by = Value.to_option (fun tac -> Proofview.tclIGNORE (thaw bt tac)) by in let ids = Value.to_list Value.to_ident ids in let cl = to_clause cl in Tac2tactics.autorewrite ~all by ids cl end -let () = define_prim1 "tac_subst" begin fun ids -> +let () = define_prim1 "tac_subst" begin fun _ ids -> let ids = Value.to_list Value.to_ident ids in Equality.subst ids end diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index a95e60bc97..5b35d53be4 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -122,62 +122,62 @@ let native where cl = let where = Option.map map_pattern_with_occs where in Tactics.reduce (CbvNative where) cl -let eval_fun red c = - Tac2core.pf_apply begin fun env sigma -> +let eval_fun bt red c = + Tac2core.pf_apply bt begin fun env sigma -> let (redfun, _) = Redexpr.reduction_of_red_expr env red in let (sigma, ans) = redfun env sigma c in Proofview.Unsafe.tclEVARS sigma >>= fun () -> Proofview.tclUNIT ans end -let eval_red c = - eval_fun (Red false) c +let eval_red bt c = + eval_fun bt (Red false) c -let eval_hnf c = - eval_fun Hnf c +let eval_hnf bt c = + eval_fun bt Hnf c -let eval_simpl flags where c = +let eval_simpl bt flags where c = let where = Option.map map_pattern_with_occs where in Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> let flags = { flags with rConst } in - eval_fun (Simpl (flags, where)) c + eval_fun bt (Simpl (flags, where)) c -let eval_cbv flags c = +let eval_cbv bt flags c = Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> let flags = { flags with rConst } in - eval_fun (Cbv flags) c + eval_fun bt (Cbv flags) c -let eval_cbn flags c = +let eval_cbn bt flags c = Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> let flags = { flags with rConst } in - eval_fun (Cbn flags) c + eval_fun bt (Cbn flags) c -let eval_lazy flags c = +let eval_lazy bt flags c = Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> let flags = { flags with rConst } in - eval_fun (Lazy flags) c + eval_fun bt (Lazy flags) c -let eval_unfold occs c = +let eval_unfold bt occs c = let map (gr, occ) = get_evaluable_reference gr >>= fun gr -> Proofview.tclUNIT (occ, gr) in Proofview.Monad.List.map map occs >>= fun occs -> - eval_fun (Unfold occs) c + eval_fun bt (Unfold occs) c -let eval_fold cl c = - eval_fun (Fold cl) c +let eval_fold bt cl c = + eval_fun bt (Fold cl) c -let eval_pattern where c = +let eval_pattern bt where c = let where = List.map (fun (pat, occ) -> (occ, pat)) where in - eval_fun (Pattern where) c + eval_fun bt (Pattern where) c -let eval_vm where c = +let eval_vm bt where c = let where = Option.map map_pattern_with_occs where in - eval_fun (CbvVm where) c + eval_fun bt (CbvVm where) c -let eval_native where c = +let eval_native bt where c = let where = Option.map map_pattern_with_occs where in - eval_fun (CbvNative where) c + eval_fun bt (CbvNative where) c let on_destruction_arg tac ev arg = Proofview.Goal.enter begin fun gl -> diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli index 87489bb248..1c76bd9648 100644 --- a/src/tac2tactics.mli +++ b/src/tac2tactics.mli @@ -9,6 +9,7 @@ open Names open Locus open Globnames +open Tac2expr open EConstr open Genredexpr open Misctypes @@ -55,28 +56,28 @@ val vm : (Pattern.constr_pattern * occurrences_expr) option -> clause -> unit ta val native : (Pattern.constr_pattern * occurrences_expr) option -> clause -> unit tactic -val eval_red : constr -> constr tactic +val eval_red : backtrace -> constr -> constr tactic -val eval_hnf : constr -> constr tactic +val eval_hnf : backtrace -> constr -> constr tactic -val eval_simpl : global_reference glob_red_flag -> +val eval_simpl : backtrace -> global_reference glob_red_flag -> (Pattern.constr_pattern * occurrences_expr) option -> constr -> constr tactic -val eval_cbv : global_reference glob_red_flag -> constr -> constr tactic +val eval_cbv : backtrace -> global_reference glob_red_flag -> constr -> constr tactic -val eval_cbn : global_reference glob_red_flag -> constr -> constr tactic +val eval_cbn : backtrace -> global_reference glob_red_flag -> constr -> constr tactic -val eval_lazy : global_reference glob_red_flag -> constr -> constr tactic +val eval_lazy : backtrace -> global_reference glob_red_flag -> constr -> constr tactic -val eval_unfold : (global_reference * occurrences_expr) list -> constr -> constr tactic +val eval_unfold : backtrace -> (global_reference * occurrences_expr) list -> constr -> constr tactic -val eval_fold : constr list -> constr -> constr tactic +val eval_fold : backtrace -> constr list -> constr -> constr tactic -val eval_pattern : (EConstr.t * occurrences_expr) list -> constr -> constr tactic +val eval_pattern : backtrace -> (EConstr.t * occurrences_expr) list -> constr -> constr tactic -val eval_vm : (Pattern.constr_pattern * occurrences_expr) option -> constr -> constr tactic +val eval_vm : backtrace -> (Pattern.constr_pattern * occurrences_expr) option -> constr -> constr tactic -val eval_native : (Pattern.constr_pattern * occurrences_expr) option -> constr -> constr tactic +val eval_native : backtrace -> (Pattern.constr_pattern * occurrences_expr) option -> constr -> constr tactic val discriminate : evars_flag -> destruction_arg option -> unit tactic -- cgit v1.2.3 From 567435828772e53327bacf7464291a5759c23831 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 4 Sep 2017 19:14:30 +0200 Subject: Better backtraces for a few datatypes. --- src/tac2core.ml | 8 ++++++-- src/tac2entries.ml | 9 +++++---- 2 files changed, 11 insertions(+), 6 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 17fa7c28f4..e4bd80adc8 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -930,9 +930,13 @@ let () = let _, tac = Genintern.intern Ltac_plugin.Tacarg.wit_tactic ist tac in GlbVal tac, gtypref t_unit in - let interp _ tac = + let interp ist tac = + let ist = { ist with env_ist = Id.Map.empty } in + let lfun = Tac2interp.set_env ist Id.Map.empty in + let ist = Ltac_plugin.Tacinterp.default_ist () in (** FUCK YOU API *) - (Obj.magic Ltac_plugin.Tacinterp.eval_tactic tac : unit Proofview.tactic) >>= fun () -> + let ist = { ist with API.Geninterp.lfun = (Obj.magic lfun) } in + (Obj.magic Ltac_plugin.Tacinterp.eval_tactic_ist ist tac : unit Proofview.tactic) >>= fun () -> return v_unit in let subst s tac = Genintern.substitute Ltac_plugin.Tacarg.wit_tactic s tac in diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 8d515577ec..197ec19b3a 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -756,14 +756,15 @@ let _ = Goptions.declare_bool_option { } let pr_frame = function -| FrLtac None -> str "" +| FrLtac None -> str "Call " | FrLtac (Some kn) -> - Libnames.pr_qualid (Tac2env.shortest_qualid_of_ltac (TacConstant kn)) + str "Call " ++ Libnames.pr_qualid (Tac2env.shortest_qualid_of_ltac (TacConstant kn)) | FrPrim ml -> - str "<" ++ str ml.mltac_plugin ++ str ":" ++ str ml.mltac_tactic ++ str ">" + str "Prim <" ++ str ml.mltac_plugin ++ str ":" ++ str ml.mltac_tactic ++ str ">" | FrExtn (tag, arg) -> let obj = Tac2env.interp_ml_object tag in - obj.Tac2env.ml_print (Global.env ()) arg + str "Extn " ++ str (Tac2dyn.Arg.repr tag) ++ str ":" ++ spc () ++ + obj.Tac2env.ml_print (Global.env ()) arg let () = register_handler begin function | Tac2interp.LtacError (kn, _, bt) -> -- cgit v1.2.3 From dd5ad74b19530568159606828c8542ac298be29d Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 4 Sep 2017 20:49:00 +0200 Subject: Implementing the non-strict mode. --- doc/ltac2.md | 24 ++++++++++++++++++++++++ src/tac2entries.ml | 6 +++--- src/tac2intern.ml | 20 +++++++++++++++++--- src/tac2intern.mli | 2 +- tests/compat.v | 4 +--- 5 files changed, 46 insertions(+), 10 deletions(-) diff --git a/doc/ltac2.md b/doc/ltac2.md index c1216d8f89..d7c8719a14 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -440,6 +440,8 @@ TERM ::= QUOTNAME := IDENT ``` +### Built-in quotations + The current implementation recognizes the following built-in quotations: - "ident", which parses identifiers (type `Init.ident`). - "constr", which parses Coq terms and produces an-evar free term at runtime @@ -457,6 +459,28 @@ The following syntactic sugar is provided for two common cases. - `@id` is the same as ident:(id) - `'t` is the same as open_constr:(t) +### Strict vs. non-strict mode + +Depending on the context, quotations producing terms (i.e. `constr` or +`open_constr`) are not internalized in the same way. There are two possible +modes, respectively called the *strict* and the *non-strict* mode. + +- In strict mode, all simple identifiers appearing in a term quotation are +required to be resolvable statically. That is, they must be the short name of +a declaration which is defined globally, excluding section variables and +hypotheses. If this doesn't hold, internalization will fail. To work around +this error, one has to specifically use the `&` notation. +- In non-strict mode, any simple identifier appearing in a term quotation which +is not bound in the global context is turned into a dynamic reference to a +hypothesis. That is to say, internalization will succeed, but the evaluation +of the term at runtime will fail if there is no such variable in the dynamic +context. + +Strict mode is enforced by default, e.g. for all Ltac2 definitions. Non-strict +mode is only set when evaluating Ltac2 snippets in interactive proof mode. The +rationale is that it is cumbersome to explicitly add `&` interactively, while it +is expected that global tactics enforce more invariants on their code. + ## Term Antiquotations ### Syntax diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 197ec19b3a..7a900e8bf0 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -325,7 +325,7 @@ let register_ltac ?(local = false) ?(mut = false) isrec tactics = if isrec then inline_rec_tactic tactics else tactics in let map ((loc, id), e) = - let (e, t) = intern e in + let (e, t) = intern ~strict:true e in let () = if not (is_value e) then user_err ?loc (str "Tactic definition must be a syntactical value") @@ -717,7 +717,7 @@ let register_redefinition ?(local = false) (loc, qid) e = if not (data.Tac2env.gdata_mutable) then user_err ?loc (str "The tactic " ++ pr_qualid qid ++ str " is not declared as mutable") in - let (e, t) = intern e in + let (e, t) = intern ~strict:true e in let () = if not (is_value e) then user_err ?loc (str "Tactic definition must be a syntactical value") @@ -826,7 +826,7 @@ let solve default tac = let call ~default e = let loc = loc_of_tacexpr e in - let (e, t) = intern e in + let (e, t) = intern ~strict:false e in let () = check_unit ?loc t in let tac = Tac2interp.interp Tac2interp.empty_environment e in solve default (Proofview.tclIGNORE tac) diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 1dba57c4c1..d1a3e13cdb 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -131,6 +131,8 @@ type environment = { (** Accept unbound type variables *) env_rec : (KerName.t * int) Id.Map.t; (** Recursive type definitions *) + env_str : bool; + (** True iff in strict mode *) } let empty_env () = { @@ -139,6 +141,7 @@ let empty_env () = { env_als = ref Id.Map.empty; env_opn = true; env_rec = Id.Map.empty; + env_str = true; } let env_name env = @@ -780,7 +783,13 @@ let rec intern_rec env (loc, e) = match e with let genv = Global.env_of_context Environ.empty_named_context_val in let ist = empty_glob_sign genv in let ist = { ist with extra = Store.set ist.extra ltac2_env env } in - let arg, tpe = Flags.with_option Ltac_plugin.Tacintern.strict_check (fun () -> obj.ml_intern self ist arg) () in + let arg, tpe = + if env.env_str then + let arg () = obj.ml_intern self ist arg in + Flags.with_option Ltac_plugin.Tacintern.strict_check arg () + else + obj.ml_intern self ist arg + in let e = match arg with | GlbVal arg -> GTacExt (tag, arg) | GlbTacexpr e -> e @@ -1121,8 +1130,9 @@ let normalize env (count, vars) (t : UF.elt glb_typexpr) = in subst_type subst t -let intern e = +let intern ~strict e = let env = empty_env () in + let env = if strict then env else { env with env_str = false } in let (e, t) = intern_rec env e in let count = ref 0 in let vars = ref UF.Map.empty in @@ -1487,7 +1497,11 @@ let () = let open Genintern in let intern ist tac = let env = match Genintern.Store.get ist.extra ltac2_env with - | None -> empty_env () + | None -> + (** Only happens when Ltac2 is called from a constr or ltac1 quotation *) + let env = empty_env () in + if !Ltac_plugin.Tacintern.strict_check then env + else { env with env_str = false } | Some env -> env in let loc = loc_of_tacexpr tac in diff --git a/src/tac2intern.mli b/src/tac2intern.mli index 045e657460..4b02f91caa 100644 --- a/src/tac2intern.mli +++ b/src/tac2intern.mli @@ -13,7 +13,7 @@ open Tac2expr val loc_of_tacexpr : raw_tacexpr -> Loc.t option val loc_of_patexpr : raw_patexpr -> Loc.t option -val intern : raw_tacexpr -> glb_tacexpr * type_scheme +val intern : strict:bool -> raw_tacexpr -> glb_tacexpr * type_scheme val intern_typedef : (KerName.t * int) Id.Map.t -> raw_quant_typedef -> glb_quant_typedef val intern_open_type : raw_typexpr -> type_scheme diff --git a/tests/compat.v b/tests/compat.v index f4e849c5de..489fa638e4 100644 --- a/tests/compat.v +++ b/tests/compat.v @@ -18,8 +18,7 @@ Qed. Goal true = false -> false = true. Proof. -(** FIXME when the non-strict mode is implemented. *) -Fail intros H; ltac1:(rewrite H); reflexivity. +intros H; ltac1:(rewrite H); reflexivity. Abort. (** Variables do not cross the compatibility layer boundary. *) @@ -57,4 +56,3 @@ Fail mytac ltac2:(fail). let t := idtac; ltac2:(fail) in mytac t. constructor. Qed. - -- cgit v1.2.3 From 0012f73a1822b97dd8bc8963bc77490cde83e89f Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 4 Sep 2017 21:30:34 +0200 Subject: More precise error messages for scope failure. --- src/tac2core.ml | 40 ++++++++++++++++++++++++++-------------- 1 file changed, 26 insertions(+), 14 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index e4bd80adc8..17fccf3ac7 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -995,7 +995,19 @@ let () = let add_scope s f = Tac2entries.register_scope (Id.of_string s) f -let scope_fail () = CErrors.user_err (str "Invalid parsing token") +let rec pr_scope = function +| SexprStr (_, s) -> qstring s +| SexprInt (_, n) -> int n +| SexprRec (_, (_, na), args) -> + let na = match na with + | None -> str "_" + | Some id -> Id.print id + in + na ++ str "(" ++ prlist_with_sep (fun () -> str ", ") pr_scope args ++ str ")" + +let scope_fail s args = + let args = str "(" ++ prlist_with_sep (fun () -> str ", ") pr_scope args ++ str ")" in + CErrors.user_err (str "Invalid arguments " ++ args ++ str " in scope " ++ str s) let q_unit = Loc.tag @@ CTacCst (AbsKn (Tuple 0)) @@ -1010,7 +1022,7 @@ let add_generic_scope s entry arg = let scope = Extend.Aentry entry in let act x = Loc.tag @@ CTacExt (arg, x) in Tac2entries.ScopeRule (scope, act) - | _ -> scope_fail () + | arg -> scope_fail s arg in add_scope s parse @@ -1018,14 +1030,14 @@ let () = add_scope "keyword" begin function | [SexprStr (loc, s)] -> let scope = Extend.Atoken (Tok.KEYWORD s) in Tac2entries.ScopeRule (scope, (fun _ -> q_unit)) -| _ -> scope_fail () +| arg -> scope_fail "keyword" arg end let () = add_scope "terminal" begin function | [SexprStr (loc, s)] -> let scope = Extend.Atoken (CLexer.terminal s) in Tac2entries.ScopeRule (scope, (fun _ -> q_unit)) -| _ -> scope_fail () +| arg -> scope_fail "terminal" arg end let () = add_scope "list0" begin function @@ -1040,7 +1052,7 @@ let () = add_scope "list0" begin function let scope = Extend.Alist0sep (scope, sep) in let act l = Tac2quote.of_list act l in Tac2entries.ScopeRule (scope, act) -| _ -> scope_fail () +| arg -> scope_fail "list0" arg end let () = add_scope "list1" begin function @@ -1055,7 +1067,7 @@ let () = add_scope "list1" begin function let scope = Extend.Alist1sep (scope, sep) in let act l = Tac2quote.of_list act l in Tac2entries.ScopeRule (scope, act) -| _ -> scope_fail () +| arg -> scope_fail "list1" arg end let () = add_scope "opt" begin function @@ -1069,7 +1081,7 @@ let () = add_scope "opt" begin function Loc.tag @@ CTacApp (Loc.tag @@ CTacCst (AbsKn (Other Core.c_some)), [act x]) in Tac2entries.ScopeRule (scope, act) -| _ -> scope_fail () +| arg -> scope_fail "opt" arg end let () = add_scope "self" begin function @@ -1077,7 +1089,7 @@ let () = add_scope "self" begin function let scope = Extend.Aself in let act tac = tac in Tac2entries.ScopeRule (scope, act) -| _ -> scope_fail () +| arg -> scope_fail "self" arg end let () = add_scope "next" begin function @@ -1085,7 +1097,7 @@ let () = add_scope "next" begin function let scope = Extend.Anext in let act tac = tac in Tac2entries.ScopeRule (scope, act) -| _ -> scope_fail () +| arg -> scope_fail "next" arg end let () = add_scope "tactic" begin function @@ -1094,12 +1106,12 @@ let () = add_scope "tactic" begin function let scope = Extend.Aentryl (tac2expr, 5) in let act tac = tac in Tac2entries.ScopeRule (scope, act) -| [SexprInt (loc, n)] -> - let () = if n < 0 || n > 6 then scope_fail () in +| [SexprInt (loc, n)] as arg -> + let () = if n < 0 || n > 6 then scope_fail "tactic" arg in let scope = Extend.Aentryl (tac2expr, n) in let act tac = tac in Tac2entries.ScopeRule (scope, act) -| _ -> scope_fail () +| arg -> scope_fail "tactic" arg end let () = add_scope "thunk" begin function @@ -1107,13 +1119,13 @@ let () = add_scope "thunk" begin function let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in let act e = rthunk (act e) in Tac2entries.ScopeRule (scope, act) -| _ -> scope_fail () +| arg -> scope_fail "thunk" arg end let add_expr_scope name entry f = add_scope name begin function | [] -> Tac2entries.ScopeRule (Extend.Aentry entry, f) - | _ -> scope_fail () + | arg -> scope_fail name arg end let () = add_expr_scope "ident" q_ident (fun id -> Tac2quote.of_anti Tac2quote.of_ident id) -- cgit v1.2.3 From 01a3776cb801ed6cbeba895d04f75e62fd6f091a Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 4 Sep 2017 21:55:51 +0200 Subject: More notations for primitive tactics. --- src/g_ltac2.ml4 | 11 +++++++---- src/tac2core.ml | 1 + src/tac2entries.ml | 1 + src/tac2entries.mli | 1 + src/tac2qexpr.mli | 4 +++- src/tac2quote.ml | 4 ++-- src/tac2quote.mli | 2 ++ theories/Notations.v | 23 +++++++++++++++++++++++ 8 files changed, 40 insertions(+), 7 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 5d5bc6b941..be7f830605 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -368,7 +368,7 @@ let loc_of_ne_list l = Loc.merge_opt (fst (List.hd l)) (fst (List.last l)) GEXTEND Gram GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause q_rewriting q_clause q_dispatch q_occurrences q_strategy_flag - q_reference q_with_bindings q_constr_matching; + q_destruction_arg q_reference q_with_bindings q_constr_matching; anti: [ [ "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] ; @@ -500,11 +500,14 @@ GEXTEND Gram [ [ c = Constr.constr; l = with_bindings -> Loc.tag ~loc:!@loc @@ (c, l) ] ] ; destruction_arg: - [ [ n = lnatural -> QElimOnAnonHyp n - | id = lident -> QElimOnIdent id - | c = constr_with_bindings -> QElimOnConstr c + [ [ n = lnatural -> Loc.tag ~loc:!@loc @@ QElimOnAnonHyp n + | id = lident -> Loc.tag ~loc:!@loc @@ QElimOnIdent id + | c = constr_with_bindings -> Loc.tag ~loc:!@loc @@ QElimOnConstr c ] ] ; + q_destruction_arg: + [ [ arg = destruction_arg -> arg ] ] + ; as_or_and_ipat: [ [ "as"; ipat = or_and_intropattern -> Some ipat | -> None diff --git a/src/tac2core.ml b/src/tac2core.ml index 17fccf3ac7..867c9ae806 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -1133,6 +1133,7 @@ let () = add_expr_scope "bindings" q_bindings Tac2quote.of_bindings let () = add_expr_scope "with_bindings" q_with_bindings Tac2quote.of_bindings let () = add_expr_scope "intropattern" q_intropattern Tac2quote.of_intro_pattern let () = add_expr_scope "intropatterns" q_intropatterns Tac2quote.of_intro_patterns +let () = add_expr_scope "destruction_arg" q_destruction_arg Tac2quote.of_destruction_arg let () = add_expr_scope "induction_clause" q_induction_clause Tac2quote.of_induction_clause let () = add_expr_scope "rewriting" q_rewriting Tac2quote.of_rewriting let () = add_expr_scope "clause" q_clause Tac2quote.of_clause diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 7a900e8bf0..afbbcfc15e 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -29,6 +29,7 @@ let q_bindings = Pcoq.Gram.entry_create "tactic:q_bindings" let q_with_bindings = Pcoq.Gram.entry_create "tactic:q_with_bindings" let q_intropattern = Pcoq.Gram.entry_create "tactic:q_intropattern" let q_intropatterns = Pcoq.Gram.entry_create "tactic:q_intropatterns" +let q_destruction_arg = Pcoq.Gram.entry_create "tactic:q_destruction_arg" let q_induction_clause = Pcoq.Gram.entry_create "tactic:q_induction_clause" let q_rewriting = Pcoq.Gram.entry_create "tactic:q_rewriting" let q_clause = Pcoq.Gram.entry_create "tactic:q_clause" diff --git a/src/tac2entries.mli b/src/tac2entries.mli index dde877666a..581d04d8cc 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -64,6 +64,7 @@ val q_bindings : bindings Pcoq.Gram.entry val q_with_bindings : bindings Pcoq.Gram.entry val q_intropattern : intro_pattern Pcoq.Gram.entry val q_intropatterns : intro_pattern list located Pcoq.Gram.entry +val q_destruction_arg : destruction_arg Pcoq.Gram.entry val q_induction_clause : induction_clause Pcoq.Gram.entry val q_rewriting : rewriting Pcoq.Gram.entry val q_clause : clause Pcoq.Gram.entry diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index a284c1e756..577fe8edfe 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -68,11 +68,13 @@ type clause = clause_r located type constr_with_bindings = (Constrexpr.constr_expr * bindings) located -type destruction_arg = +type destruction_arg_r = | QElimOnConstr of constr_with_bindings | QElimOnIdent of Id.t located | QElimOnAnonHyp of int located +type destruction_arg = destruction_arg_r located + type induction_clause_r = { indcl_arg : destruction_arg; indcl_eqn : intro_pattern_naming option; diff --git a/src/tac2quote.ml b/src/tac2quote.ml index d38d7414fe..ed4cef2e2a 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -191,7 +191,7 @@ let of_clause (loc, cl) = std_proj "on_concl", concl; ]) -let of_destruction_arg ?loc = function +let of_destruction_arg (loc, arg) = match arg with | QElimOnConstr c -> let arg = thunk (of_constr_with_bindings c) in std_constructor ?loc "ElimOnConstr" [arg] @@ -199,7 +199,7 @@ let of_destruction_arg ?loc = function | QElimOnAnonHyp n -> std_constructor ?loc "ElimOnAnonHyp" [of_int n] let of_induction_clause (loc, cl) = - let arg = of_destruction_arg ?loc cl.indcl_arg in + let arg = of_destruction_arg cl.indcl_arg in let eqn = of_option ?loc of_intro_pattern_naming cl.indcl_eqn in let as_ = of_option ?loc of_or_and_intro_pattern cl.indcl_as in let in_ = of_option ?loc of_clause cl.indcl_in in diff --git a/src/tac2quote.mli b/src/tac2quote.mli index c3374ac24e..875230b7e3 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -46,6 +46,8 @@ val of_intro_patterns : intro_pattern list located -> raw_tacexpr val of_clause : clause -> raw_tacexpr +val of_destruction_arg : destruction_arg -> raw_tacexpr + val of_induction_clause : induction_clause -> raw_tacexpr val of_rewriting : rewriting -> raw_tacexpr diff --git a/theories/Notations.v b/theories/Notations.v index 5ed47336ad..ad89bc5cfc 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -383,6 +383,12 @@ Ltac2 Notation "eexact" c(thunk(open_constr)) := exact0 true c. Ltac2 Notation reflexivity := Std.reflexivity (). +Ltac2 symmetry0 cl := + Std.symmetry (default_on_concl cl). + +Ltac2 Notation "symmetry" cl(opt(clause)) := symmetry0 cl. +Ltac2 Notation symmetry := symmetry. + Ltac2 Notation assumption := Std.assumption (). Ltac2 Notation etransitivity := Std.etransitivity (). @@ -412,3 +418,20 @@ end. Ltac2 Notation "subst" ids(list0(ident)) := subst0 ids. Ltac2 Notation subst := subst. + +Ltac2 Notation "discriminate" arg(opt(destruction_arg)) := + Std.discriminate false arg. + +Ltac2 Notation "ediscriminate" arg(opt(destruction_arg)) := + Std.discriminate true arg. + +Ltac2 Notation "injection" arg(opt(destruction_arg)) ipat(opt(seq("as", intropatterns))):= + Std.injection false ipat arg. + +Ltac2 Notation "einjection" arg(opt(destruction_arg)) ipat(opt(seq("as", intropatterns))):= + Std.injection true ipat arg. + +(** Congruence *) + +Ltac2 f_equal0 () := ltac1:(f_equal). +Ltac2 Notation f_equal := f_equal0 (). -- cgit v1.2.3 From ac3bb720d4ec04aa670845352df1d8b8885f865e Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 4 Sep 2017 22:42:33 +0200 Subject: Notation for "clear - ids". --- theories/Notations.v | 1 + 1 file changed, 1 insertion(+) diff --git a/theories/Notations.v b/theories/Notations.v index ad89bc5cfc..136ca871c2 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -401,6 +401,7 @@ Ltac2 clear0 ids := match ids with end. Ltac2 Notation "clear" ids(list0(ident)) := clear0 ids. +Ltac2 Notation "clear" "-" ids(list1(ident)) := Std.keep ids. Ltac2 Notation clear := clear. Ltac2 Notation refine := Control.refine. -- cgit v1.2.3 From ada4c3aadb5c6b1870c2bf962ef9e1b07cc4bb05 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 4 Sep 2017 23:54:30 +0200 Subject: ML bindings of auto-related tactics. --- src/tac2stdlib.ml | 64 +++++++++++++++++++++++++++++++++++++++++++++++++++++ src/tac2tactics.ml | 30 +++++++++++++++++++++++++ src/tac2tactics.mli | 15 +++++++++++++ theories/Std.v | 16 ++++++++++++++ 4 files changed, 125 insertions(+) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 6dcb3f15fb..ed0e6aafd3 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -159,6 +159,17 @@ let to_rewriting = function (orient, repeat, c) | _ -> assert false +let to_debug = function +| ValInt 0 -> Hints.Off +| ValInt 1 -> Hints.Info +| ValInt 2 -> Hints.Debug +| _ -> assert false + +let to_strategy = function +| ValInt 0 -> Class_tactics.Bfs +| ValInt 1 -> Class_tactics.Dfs +| _ -> assert false + (** Standard tactics sharing their implementation with Ltac1 *) let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } @@ -200,6 +211,13 @@ let define_prim4 name tac = in Tac2env.define_primitive (pname name) tac +let define_prim5 name tac = + let tac bt arg = match arg with + | [x; y; z; u; v] -> lift (tac bt x y z u v) + | _ -> assert false + in + Tac2env.define_primitive (pname name) tac + (** Tactics from Tacexpr *) let () = define_prim2 "tac_intros" begin fun _ ev ipat -> @@ -608,3 +626,49 @@ let () = define_prim1 "tac_subst" begin fun _ ids -> end let () = define_prim0 "tac_substall" (return () >>= fun () -> Equality.subst_all ()) + +(** Auto *) + +let () = define_prim3 "tac_trivial" begin fun bt dbg lems dbs -> + let dbg = to_debug dbg in + let map c = thaw bt c >>= fun c -> return (Value.to_constr c) in + let lems = Value.to_list map lems in + let dbs = Value.to_option (fun l -> Value.to_list Value.to_ident l) dbs in + Tac2tactics.trivial dbg lems dbs +end + +let () = define_prim5 "tac_eauto" begin fun bt dbg n p lems dbs -> + let dbg = to_debug dbg in + let n = Value.to_option Value.to_int n in + let p = Value.to_option Value.to_int p in + let map c = thaw bt c >>= fun c -> return (Value.to_constr c) in + let lems = Value.to_list map lems in + let dbs = Value.to_option (fun l -> Value.to_list Value.to_ident l) dbs in + Tac2tactics.eauto dbg n p lems dbs +end + +let () = define_prim4 "tac_auto" begin fun bt dbg n lems dbs -> + let dbg = to_debug dbg in + let n = Value.to_option Value.to_int n in + let map c = thaw bt c >>= fun c -> return (Value.to_constr c) in + let lems = Value.to_list map lems in + let dbs = Value.to_option (fun l -> Value.to_list Value.to_ident l) dbs in + Tac2tactics.auto dbg n lems dbs +end + +let () = define_prim4 "tac_newauto" begin fun bt dbg n lems dbs -> + let dbg = to_debug dbg in + let n = Value.to_option Value.to_int n in + let map c = thaw bt c >>= fun c -> return (Value.to_constr c) in + let lems = Value.to_list map lems in + let dbs = Value.to_option (fun l -> Value.to_list Value.to_ident l) dbs in + Tac2tactics.new_auto dbg n lems dbs +end + +let () = define_prim4 "tac_typeclasses_eauto" begin fun bt b str n dbs -> + let b = Value.to_bool b in + let str = to_strategy str in + let n = Value.to_option Value.to_int n in + let dbs = Value.to_list Value.to_ident dbs in + Tac2tactics.typeclasses_eauto b str n dbs +end diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index 5b35d53be4..b35e26c89e 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -216,3 +216,33 @@ let autorewrite ~all by ids cl = match by with | None -> Autorewrite.auto_multi_rewrite ?conds ids cl | Some by -> Autorewrite.auto_multi_rewrite_with ?conds by ids cl + +(** Auto *) + +let trivial debug lems dbs = + let lems = List.map delayed_of_tactic lems in + let dbs = Option.map (fun l -> List.map Id.to_string l) dbs in + Auto.h_trivial ~debug lems dbs + +let auto debug n lems dbs = + let lems = List.map delayed_of_tactic lems in + let dbs = Option.map (fun l -> List.map Id.to_string l) dbs in + Auto.h_auto ~debug n lems dbs + +let new_auto debug n lems dbs = + let make_depth n = snd (Eauto.make_dimension n None) in + let lems = List.map delayed_of_tactic lems in + match dbs with + | None -> Auto.new_full_auto ~debug (make_depth n) lems + | Some dbs -> + let dbs = List.map Id.to_string dbs in + Auto.new_auto ~debug (make_depth n) lems dbs + +let eauto debug n p lems dbs = + let lems = List.map delayed_of_tactic lems in + let dbs = Option.map (fun l -> List.map Id.to_string l) dbs in + Eauto.gen_eauto (Eauto.make_dimension n p) lems dbs + +let typeclasses_eauto only_classes strategy depth dbs = + let dbs = List.map Id.to_string dbs in + Class_tactics.typeclasses_eauto ~only_classes ~strategy ~depth dbs diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli index 1c76bd9648..4369919d31 100644 --- a/src/tac2tactics.mli +++ b/src/tac2tactics.mli @@ -84,3 +84,18 @@ val discriminate : evars_flag -> destruction_arg option -> unit tactic val injection : evars_flag -> intro_pattern list option -> destruction_arg option -> unit tactic val autorewrite : all:bool -> unit tactic option -> Id.t list -> clause -> unit tactic + +val trivial : Hints.debug -> constr tactic list -> Id.t list option -> + unit Proofview.tactic + +val auto : Hints.debug -> int option -> constr tactic list -> + Id.t list option -> unit Proofview.tactic + +val new_auto : Hints.debug -> int option -> constr tactic list -> + Id.t list option -> unit Proofview.tactic + +val eauto : Hints.debug -> int option -> int option -> constr tactic list -> + Id.t list option -> unit Proofview.tactic + +val typeclasses_eauto : bool -> Class_tactics.search_strategy -> int option -> + Id.t list -> unit Proofview.tactic diff --git a/theories/Std.v b/theories/Std.v index 2fa2c34da6..b667258aa2 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -211,3 +211,19 @@ Ltac2 @ external autorewrite : bool -> (unit -> unit) option -> ident list -> cl Ltac2 @ external subst : ident list -> unit := "ltac2" "tac_subst". Ltac2 @ external subst_all : unit -> unit := "ltac2" "tac_substall". + +(** auto *) + +Ltac2 Type debug := [ Off | Info | Debug ]. + +Ltac2 Type strategy := [ BFS | DFS ]. + +Ltac2 @ external trivial : debug -> (unit -> constr) list -> ident list option -> unit := "ltac2" "tac_trivial". + +Ltac2 @ external auto : debug -> int option -> (unit -> constr) list -> ident list option -> unit := "ltac2" "tac_auto". + +Ltac2 @ external new_auto : debug -> int option -> (unit -> constr) list -> ident list option -> unit := "ltac2" "tac_newauto". + +Ltac2 @ external eauto : debug -> int option -> int option -> (unit -> constr) list -> ident list option -> unit := "ltac2" "tac_eauto". + +Ltac2 @ external typeclasses_eauto : bool -> strategy -> int option -> ident list -> unit := "ltac2" "tac_typeclasses_eauto". -- cgit v1.2.3 From ebe95a28cf012aff33eb5ce167be6520e6643cfd Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 5 Sep 2017 00:35:49 +0200 Subject: More static invariants for typeclass_eauto. --- src/tac2stdlib.ml | 7 +++---- src/tac2tactics.ml | 10 ++++++++-- src/tac2tactics.mli | 4 ++-- theories/Std.v | 2 +- 4 files changed, 14 insertions(+), 9 deletions(-) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index ed0e6aafd3..b64aac3559 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -665,10 +665,9 @@ let () = define_prim4 "tac_newauto" begin fun bt dbg n lems dbs -> Tac2tactics.new_auto dbg n lems dbs end -let () = define_prim4 "tac_typeclasses_eauto" begin fun bt b str n dbs -> - let b = Value.to_bool b in +let () = define_prim3 "tac_typeclasses_eauto" begin fun bt str n dbs -> let str = to_strategy str in let n = Value.to_option Value.to_int n in - let dbs = Value.to_list Value.to_ident dbs in - Tac2tactics.typeclasses_eauto b str n dbs + let dbs = Value.to_option (fun l -> Value.to_list Value.to_ident l) dbs in + Tac2tactics.typeclasses_eauto str n dbs end diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index b35e26c89e..6cf8f24f27 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -243,6 +243,12 @@ let eauto debug n p lems dbs = let dbs = Option.map (fun l -> List.map Id.to_string l) dbs in Eauto.gen_eauto (Eauto.make_dimension n p) lems dbs -let typeclasses_eauto only_classes strategy depth dbs = - let dbs = List.map Id.to_string dbs in +let typeclasses_eauto strategy depth dbs = + let only_classes, dbs = match dbs with + | None -> + true, [Hints.typeclasses_db] + | Some dbs -> + let dbs = List.map Id.to_string dbs in + false, dbs + in Class_tactics.typeclasses_eauto ~only_classes ~strategy ~depth dbs diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli index 4369919d31..e0cd77096b 100644 --- a/src/tac2tactics.mli +++ b/src/tac2tactics.mli @@ -97,5 +97,5 @@ val new_auto : Hints.debug -> int option -> constr tactic list -> val eauto : Hints.debug -> int option -> int option -> constr tactic list -> Id.t list option -> unit Proofview.tactic -val typeclasses_eauto : bool -> Class_tactics.search_strategy -> int option -> - Id.t list -> unit Proofview.tactic +val typeclasses_eauto : Class_tactics.search_strategy -> int option -> + Id.t list option -> unit Proofview.tactic diff --git a/theories/Std.v b/theories/Std.v index b667258aa2..5201fa819d 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -226,4 +226,4 @@ Ltac2 @ external new_auto : debug -> int option -> (unit -> constr) list -> iden Ltac2 @ external eauto : debug -> int option -> int option -> (unit -> constr) list -> ident list option -> unit := "ltac2" "tac_eauto". -Ltac2 @ external typeclasses_eauto : bool -> strategy -> int option -> ident list -> unit := "ltac2" "tac_typeclasses_eauto". +Ltac2 @ external typeclasses_eauto : strategy -> int option -> ident list option -> unit := "ltac2" "tac_typeclasses_eauto". -- cgit v1.2.3 From da28b6e65d9b9a74c277cb15055131c8a151bb72 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 5 Sep 2017 00:42:22 +0200 Subject: Quotations for auto-related tactics. --- src/g_ltac2.ml4 | 11 ++++++++- src/tac2core.ml | 1 + src/tac2entries.ml | 1 + src/tac2entries.mli | 1 + src/tac2qexpr.mli | 6 +++++ src/tac2quote.ml | 4 ++++ src/tac2quote.mli | 2 ++ theories/Notations.v | 66 ++++++++++++++++++++++++++++++++++++++++++++++++++++ theories/Std.v | 2 +- 9 files changed, 92 insertions(+), 2 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index be7f830605..67254d0781 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -368,7 +368,8 @@ let loc_of_ne_list l = Loc.merge_opt (fst (List.hd l)) (fst (List.last l)) GEXTEND Gram GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause q_rewriting q_clause q_dispatch q_occurrences q_strategy_flag - q_destruction_arg q_reference q_with_bindings q_constr_matching; + q_destruction_arg q_reference q_with_bindings q_constr_matching + q_hintdb; anti: [ [ "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] ; @@ -664,6 +665,14 @@ GEXTEND Gram q_strategy_flag: [ [ flag = strategy_flag -> flag ] ] ; + hintdb: + [ [ "*" -> Loc.tag ~loc:!@loc @@ QHintAll + | l = LIST1 ident_or_anti -> Loc.tag ~loc:!@loc @@ QHintDbs l + ] ] + ; + q_hintdb: + [ [ db = hintdb -> db ] ] + ; match_pattern: [ [ IDENT "context"; id = OPT Prim.ident; "["; pat = Constr.lconstr_pattern; "]" -> (Some id, pat) diff --git a/src/tac2core.ml b/src/tac2core.ml index 867c9ae806..a735dd19d9 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -1137,6 +1137,7 @@ let () = add_expr_scope "destruction_arg" q_destruction_arg Tac2quote.of_destruc let () = add_expr_scope "induction_clause" q_induction_clause Tac2quote.of_induction_clause let () = add_expr_scope "rewriting" q_rewriting Tac2quote.of_rewriting let () = add_expr_scope "clause" q_clause Tac2quote.of_clause +let () = add_expr_scope "hintdb" q_hintdb Tac2quote.of_hintdb let () = add_expr_scope "occurrences" q_occurrences Tac2quote.of_occurrences let () = add_expr_scope "dispatch" q_dispatch Tac2quote.of_dispatch let () = add_expr_scope "strategy" q_strategy_flag Tac2quote.of_strategy_flag diff --git a/src/tac2entries.ml b/src/tac2entries.ml index afbbcfc15e..121841e8dc 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -38,6 +38,7 @@ let q_occurrences = Pcoq.Gram.entry_create "tactic:q_occurrences" let q_reference = Pcoq.Gram.entry_create "tactic:q_reference" let q_strategy_flag = Pcoq.Gram.entry_create "tactic:q_strategy_flag" let q_constr_matching = Pcoq.Gram.entry_create "tactic:q_constr_matching" +let q_hintdb = Pcoq.Gram.entry_create "tactic:q_hintdb" end (** Tactic definition *) diff --git a/src/tac2entries.mli b/src/tac2entries.mli index 581d04d8cc..91e2a683d8 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -73,6 +73,7 @@ val q_occurrences : occurrences Pcoq.Gram.entry val q_reference : Libnames.reference or_anti Pcoq.Gram.entry val q_strategy_flag : strategy_flag Pcoq.Gram.entry val q_constr_matching : constr_matching Pcoq.Gram.entry +val q_hintdb : hintdb Pcoq.Gram.entry end (** {5 Hooks} *) diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index 577fe8edfe..4bbaf43d8d 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -125,3 +125,9 @@ type constr_match_branch_r = type constr_match_branch = constr_match_branch_r located type constr_matching = constr_match_branch list located + +type hintdb_r = +| QHintAll +| QHintDbs of Id.t located or_anti list + +type hintdb = hintdb_r located diff --git a/src/tac2quote.ml b/src/tac2quote.ml index ed4cef2e2a..f87985435c 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -305,6 +305,10 @@ let of_strategy_flag (loc, flag) = std_proj "rConst", of_list ?loc of_reference flag.rConst; ]) +let of_hintdb (loc, hdb) = match hdb with +| QHintAll -> of_option ?loc (fun l -> of_list (fun id -> of_anti of_ident id) l) None +| QHintDbs ids -> of_option ?loc (fun l -> of_list (fun id -> of_anti of_ident id) l) (Some ids) + let pattern_vars pat = let rec aux () accu pat = match pat.CAst.v with | Constrexpr.CPatVar id -> Id.Set.add id accu diff --git a/src/tac2quote.mli b/src/tac2quote.mli index 875230b7e3..b85f3438a3 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -54,6 +54,8 @@ val of_rewriting : rewriting -> raw_tacexpr val of_occurrences : occurrences -> raw_tacexpr +val of_hintdb : hintdb -> raw_tacexpr + val of_reference : Libnames.reference or_anti -> raw_tacexpr val of_hyp : ?loc:Loc.t -> Id.t located -> raw_tacexpr diff --git a/theories/Notations.v b/theories/Notations.v index 136ca871c2..4cb4f32682 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -432,6 +432,72 @@ Ltac2 Notation "injection" arg(opt(destruction_arg)) ipat(opt(seq("as", intropat Ltac2 Notation "einjection" arg(opt(destruction_arg)) ipat(opt(seq("as", intropatterns))):= Std.injection true ipat arg. +(** Auto *) + +Ltac2 default_db dbs := match dbs with +| None => Some [] +| Some dbs => + match dbs with + | None => None + | Some l => Some l + end +end. + +Ltac2 default_using use := match use with +| None => [] +| Some use => use +end. + +Ltac2 trivial0 use dbs := + let dbs := default_db dbs in + let use := default_using use in + Std.trivial Std.Off use dbs. + +Ltac2 Notation "trivial" + use(opt(seq("using", list1(thunk(constr), ",")))) + dbs(opt(seq("with", hintdb))) := trivial0 use dbs. + +Ltac2 Notation trivial := trivial. + +Ltac2 auto0 n use dbs := + let dbs := default_db dbs in + let use := default_using use in + Std.auto Std.Off n use dbs. + +Ltac2 Notation "auto" n(opt(tactic(0))) + use(opt(seq("using", list1(thunk(constr), ",")))) + dbs(opt(seq("with", hintdb))) := auto0 n use dbs. + +Ltac2 Notation auto := auto. + +Ltac2 new_eauto0 n use dbs := + let dbs := default_db dbs in + let use := default_using use in + Std.new_auto Std.Off n use dbs. + +Ltac2 Notation "new" "auto" n(opt(tactic(0))) + use(opt(seq("using", list1(thunk(constr), ",")))) + dbs(opt(seq("with", hintdb))) := new_eauto0 n use dbs. + +Ltac2 eauto0 n p use dbs := + let dbs := default_db dbs in + let use := default_using use in + Std.eauto Std.Off n p use dbs. + +Ltac2 Notation "eauto" n(opt(tactic(0))) p(opt(tactic(0))) + use(opt(seq("using", list1(thunk(constr), ",")))) + dbs(opt(seq("with", hintdb))) := eauto0 n p use dbs. + +Ltac2 Notation eauto := eauto. + +Ltac2 Notation "typeclasses_eauto" n(opt(tactic(0))) p(opt(tactic(0))) + dbs(opt(seq("with", list1(ident)))) := Std.typeclasses_eauto Std.DFS n dbs. + +Ltac2 Notation "typeclasses_eauto" "bfs" n(opt(tactic(0))) p(opt(tactic(0))) + dbs(opt(seq("with", list1(ident)))) := Std.typeclasses_eauto Std.BFS n dbs. + +Ltac2 Notation typeclasses_eauto := typeclasses_eauto. + (** Congruence *) Ltac2 f_equal0 () := ltac1:(f_equal). diff --git a/theories/Std.v b/theories/Std.v index 5201fa819d..79a7be1d63 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -216,7 +216,7 @@ Ltac2 @ external subst_all : unit -> unit := "ltac2" "tac_substall". Ltac2 Type debug := [ Off | Info | Debug ]. -Ltac2 Type strategy := [ BFS | DFS ]. +Ltac2 Type strategy := [ BFS | DFS ]. Ltac2 @ external trivial : debug -> (unit -> constr) list -> ident list option -> unit := "ltac2" "tac_trivial". -- cgit v1.2.3 From 9a3853fe872e200ed1b34319f6ff0f85a171a434 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 5 Sep 2017 01:07:04 +0200 Subject: Fixup grammar of typeclasses_eauto. --- theories/Notations.v | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/theories/Notations.v b/theories/Notations.v index 4cb4f32682..e7fc43c0ee 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -490,10 +490,10 @@ Ltac2 Notation "eauto" n(opt(tactic(0))) p(opt(tactic(0))) Ltac2 Notation eauto := eauto. -Ltac2 Notation "typeclasses_eauto" n(opt(tactic(0))) p(opt(tactic(0))) +Ltac2 Notation "typeclasses_eauto" n(opt(tactic(0))) dbs(opt(seq("with", list1(ident)))) := Std.typeclasses_eauto Std.DFS n dbs. -Ltac2 Notation "typeclasses_eauto" "bfs" n(opt(tactic(0))) p(opt(tactic(0))) +Ltac2 Notation "typeclasses_eauto" "bfs" n(opt(tactic(0))) dbs(opt(seq("with", list1(ident)))) := Std.typeclasses_eauto Std.BFS n dbs. Ltac2 Notation typeclasses_eauto := typeclasses_eauto. -- cgit v1.2.3 From 4a8399e62dd4bdf5876e714910dd2c7cb433dfda Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 5 Sep 2017 01:29:21 +0200 Subject: Typeclasses_eauto strategy is now optional. --- src/tac2stdlib.ml | 2 +- src/tac2tactics.ml | 2 +- src/tac2tactics.mli | 2 +- theories/Notations.v | 4 ++-- theories/Std.v | 2 +- 5 files changed, 6 insertions(+), 6 deletions(-) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index b64aac3559..a0eb0d60e5 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -666,7 +666,7 @@ let () = define_prim4 "tac_newauto" begin fun bt dbg n lems dbs -> end let () = define_prim3 "tac_typeclasses_eauto" begin fun bt str n dbs -> - let str = to_strategy str in + let str = Value.to_option to_strategy str in let n = Value.to_option Value.to_int n in let dbs = Value.to_option (fun l -> Value.to_list Value.to_ident l) dbs in Tac2tactics.typeclasses_eauto str n dbs diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index 6cf8f24f27..447f602f7a 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -251,4 +251,4 @@ let typeclasses_eauto strategy depth dbs = let dbs = List.map Id.to_string dbs in false, dbs in - Class_tactics.typeclasses_eauto ~only_classes ~strategy ~depth dbs + Class_tactics.typeclasses_eauto ~only_classes ?strategy ~depth dbs diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli index e0cd77096b..78d421303a 100644 --- a/src/tac2tactics.mli +++ b/src/tac2tactics.mli @@ -97,5 +97,5 @@ val new_auto : Hints.debug -> int option -> constr tactic list -> val eauto : Hints.debug -> int option -> int option -> constr tactic list -> Id.t list option -> unit Proofview.tactic -val typeclasses_eauto : Class_tactics.search_strategy -> int option -> +val typeclasses_eauto : Class_tactics.search_strategy option -> int option -> Id.t list option -> unit Proofview.tactic diff --git a/theories/Notations.v b/theories/Notations.v index e7fc43c0ee..5fdb4ec8af 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -491,10 +491,10 @@ Ltac2 Notation "eauto" n(opt(tactic(0))) p(opt(tactic(0))) Ltac2 Notation eauto := eauto. Ltac2 Notation "typeclasses_eauto" n(opt(tactic(0))) - dbs(opt(seq("with", list1(ident)))) := Std.typeclasses_eauto Std.DFS n dbs. + dbs(opt(seq("with", list1(ident)))) := Std.typeclasses_eauto None n dbs. Ltac2 Notation "typeclasses_eauto" "bfs" n(opt(tactic(0))) - dbs(opt(seq("with", list1(ident)))) := Std.typeclasses_eauto Std.BFS n dbs. + dbs(opt(seq("with", list1(ident)))) := Std.typeclasses_eauto (Some Std.BFS) n dbs. Ltac2 Notation typeclasses_eauto := typeclasses_eauto. diff --git a/theories/Std.v b/theories/Std.v index 79a7be1d63..b63c2eaa41 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -226,4 +226,4 @@ Ltac2 @ external new_auto : debug -> int option -> (unit -> constr) list -> iden Ltac2 @ external eauto : debug -> int option -> int option -> (unit -> constr) list -> ident list option -> unit := "ltac2" "tac_eauto". -Ltac2 @ external typeclasses_eauto : strategy -> int option -> ident list option -> unit := "ltac2" "tac_typeclasses_eauto". +Ltac2 @ external typeclasses_eauto : strategy option -> int option -> ident list option -> unit := "ltac2" "tac_typeclasses_eauto". -- cgit v1.2.3 From ca40f89c7be05253ea04585ac9ce068aa4744ae9 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 5 Sep 2017 01:39:51 +0200 Subject: A few more notations. --- theories/Notations.v | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/theories/Notations.v b/theories/Notations.v index 5fdb4ec8af..75abb2f1cd 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -389,6 +389,8 @@ Ltac2 symmetry0 cl := Ltac2 Notation "symmetry" cl(opt(clause)) := symmetry0 cl. Ltac2 Notation symmetry := symmetry. +Ltac2 Notation "revert" ids(list1(ident)) := Std.revert ids. + Ltac2 Notation assumption := Std.assumption (). Ltac2 Notation etransitivity := Std.etransitivity (). @@ -422,9 +424,11 @@ Ltac2 Notation subst := subst. Ltac2 Notation "discriminate" arg(opt(destruction_arg)) := Std.discriminate false arg. +Ltac2 Notation discriminate := discriminate. Ltac2 Notation "ediscriminate" arg(opt(destruction_arg)) := Std.discriminate true arg. +Ltac2 Notation ediscriminate := ediscriminate. Ltac2 Notation "injection" arg(opt(destruction_arg)) ipat(opt(seq("as", intropatterns))):= Std.injection false ipat arg. -- cgit v1.2.3 From 3e71c616fdafd86652bf9e14505ae1379a6f37bc Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 5 Sep 2017 12:23:29 +0200 Subject: Binding the inversion family of tactics. --- src/tac2stdlib.ml | 26 ++++++++++++++++++++------ src/tac2tactics.ml | 29 +++++++++++++++++++++++++++++ src/tac2tactics.mli | 2 ++ theories/Notations.v | 18 ++++++++++++++++++ theories/Std.v | 8 ++++++++ 5 files changed, 77 insertions(+), 6 deletions(-) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index a0eb0d60e5..03141805ef 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -128,9 +128,9 @@ let to_destruction_arg = function | ValBlk (0, [| c |]) -> (** FIXME: lost backtrace *) let c = thaw [] c >>= fun c -> return (to_constr_with_bindings c) in - ElimOnConstr c -| ValBlk (1, [| id |]) -> ElimOnIdent (Loc.tag (Value.to_ident id)) -| ValBlk (2, [| n |]) -> ElimOnAnonHyp (Value.to_int n) + None, ElimOnConstr c +| ValBlk (1, [| id |]) -> None, ElimOnIdent (Loc.tag (Value.to_ident id)) +| ValBlk (2, [| n |]) -> None, ElimOnAnonHyp (Value.to_int n) | _ -> assert false let to_induction_clause = function @@ -139,7 +139,7 @@ let to_induction_clause = function let eqn = Value.to_option (fun p -> Loc.tag (to_intro_pattern_naming p)) eqn in let as_ = Value.to_option (fun p -> Loc.tag (to_or_and_intro_pattern p)) as_ in let in_ = Value.to_option to_clause in_ in - ((None, arg), eqn, as_, in_) + (arg, eqn, as_, in_) | _ -> assert false @@ -170,6 +170,12 @@ let to_strategy = function | ValInt 1 -> Class_tactics.Dfs | _ -> assert false +let to_inversion_kind = function +| ValInt 0 -> Misctypes.SimpleInversion +| ValInt 1 -> Misctypes.FullInversion +| ValInt 2 -> Misctypes.FullInversionClear +| _ -> assert false + (** Standard tactics sharing their implementation with Ltac1 *) let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } @@ -470,6 +476,14 @@ let () = define_prim4 "tac_rewrite" begin fun bt ev rw cl by -> Tac2tactics.rewrite ev rw cl by end +let () = define_prim4 "tac_inversion" begin fun bt knd arg pat ids -> + let knd = to_inversion_kind knd in + let arg = to_destruction_arg arg in + let pat = Value.to_option (fun ipat -> Loc.tag (to_intro_pattern ipat)) pat in + let ids = Value.to_option (fun l -> Value.to_list Value.to_ident l) ids in + Tac2tactics.inversion knd arg pat ids +end + (** Tactics from coretactics *) let () = define_prim0 "tac_reflexivity" Tactics.intros_reflexivity @@ -592,14 +606,14 @@ end let () = define_prim2 "tac_discriminate" begin fun _ ev arg -> let ev = Value.to_bool ev in - let arg = Value.to_option (fun arg -> None, to_destruction_arg arg) arg in + let arg = Value.to_option to_destruction_arg arg in Tac2tactics.discriminate ev arg end let () = define_prim3 "tac_injection" begin fun _ ev ipat arg -> let ev = Value.to_bool ev in let ipat = Value.to_option to_intro_patterns ipat in - let arg = Value.to_option (fun arg -> None, to_destruction_arg arg) arg in + let arg = Value.to_option to_destruction_arg arg in Tac2tactics.injection ev ipat arg end diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index 447f602f7a..b069e57235 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -252,3 +252,32 @@ let typeclasses_eauto strategy depth dbs = false, dbs in Class_tactics.typeclasses_eauto ~only_classes ?strategy ~depth dbs + +(** Inversion *) + +let inversion knd arg pat ids = + let ids = match ids with + | None -> [] + | Some l -> l + in + begin match pat with + | None -> Proofview.tclUNIT None + | Some (_, IntroAction (IntroOrAndPattern p)) -> + Proofview.tclUNIT (Some (Loc.tag p)) + | Some _ -> + Tacticals.New.tclZEROMSG (str "Inversion only accept disjunctive patterns") + end >>= fun pat -> + let inversion _ arg = + begin match arg with + | None -> assert false + | Some (_, ElimOnAnonHyp n) -> + Inv.inv_clause knd pat ids (AnonHyp n) + | Some (_, ElimOnIdent (_, id)) -> + Inv.inv_clause knd pat ids (NamedHyp id) + | Some (_, ElimOnConstr c) -> + let anon = Loc.tag @@ IntroNaming IntroAnonymous in + Tactics.specialize c (Some anon) >>= fun () -> + Tacticals.New.onLastHypId (fun id -> Inv.inv_clause knd pat ids (NamedHyp id)) + end + in + on_destruction_arg inversion true (Some arg) diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli index 78d421303a..8b466cd529 100644 --- a/src/tac2tactics.mli +++ b/src/tac2tactics.mli @@ -99,3 +99,5 @@ val eauto : Hints.debug -> int option -> int option -> constr tactic list -> val typeclasses_eauto : Class_tactics.search_strategy option -> int option -> Id.t list option -> unit Proofview.tactic + +val inversion : inversion_kind -> destruction_arg -> intro_pattern option -> Id.t list option -> unit tactic diff --git a/theories/Notations.v b/theories/Notations.v index 75abb2f1cd..8a3e769d12 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -298,6 +298,24 @@ Ltac2 Notation "edestruct" use(thunk(opt(seq("using", constr, with_bindings)))) := destruct0 true ic use. +Ltac2 Notation "simple" "inversion" + arg(destruction_arg) + pat(opt(seq("as", intropattern))) + ids(opt(seq("in", list1(ident)))) := + Std.inversion Std.SimpleInversion arg pat ids. + +Ltac2 Notation "inversion" + arg(destruction_arg) + pat(opt(seq("as", intropattern))) + ids(opt(seq("in", list1(ident)))) := + Std.inversion Std.FullInversion arg pat ids. + +Ltac2 Notation "inversion_clear" + arg(destruction_arg) + pat(opt(seq("as", intropattern))) + ids(opt(seq("in", list1(ident)))) := + Std.inversion Std.FullInversionClear arg pat ids. + Ltac2 default_on_concl cl := match cl with | None => { Std.on_hyps := Some []; Std.on_concl := Std.AllOccurrences } diff --git a/theories/Std.v b/theories/Std.v index b63c2eaa41..02bc4ff450 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -106,6 +106,12 @@ Ltac2 Type rewriting := { Ltac2 Type evar_flag := bool. Ltac2 Type advanced_flag := bool. +Ltac2 Type inversion_kind := [ +| SimpleInversion +| FullInversion +| FullInversionClear +]. + (** Standard, built-in tactics. See Ltac1 for documentation. *) Ltac2 @ external intros : evar_flag -> intro_pattern list -> unit := "ltac2" "tac_intros". @@ -197,6 +203,8 @@ Ltac2 @ external exact_no_check : constr -> unit := "ltac2" "tac_exactnocheck". Ltac2 @ external vm_cast_no_check : constr -> unit := "ltac2" "tac_vmcastnocheck". Ltac2 @ external native_cast_no_check : constr -> unit := "ltac2" "tac_nativecastnocheck". +Ltac2 @ external inversion : inversion_kind -> destruction_arg -> intro_pattern option -> ident list option -> unit := "ltac2" "tac_inversion". + (** coretactics *) (** extratactics *) -- cgit v1.2.3 From 2b0e0ad1062ad49c8bd7d4a7d183fe0119f81803 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 5 Sep 2017 15:39:04 +0200 Subject: Introducing quotations for move locations. --- src/g_ltac2.ml4 | 12 +++++++++++- src/tac2core.ml | 1 + src/tac2entries.ml | 1 + src/tac2entries.mli | 1 + src/tac2qexpr.mli | 8 ++++++++ src/tac2quote.ml | 6 ++++++ src/tac2quote.mli | 2 ++ theories/Std.v | 7 +++++++ 8 files changed, 37 insertions(+), 1 deletion(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 67254d0781..5c285010e9 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -369,7 +369,7 @@ GEXTEND Gram GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause q_rewriting q_clause q_dispatch q_occurrences q_strategy_flag q_destruction_arg q_reference q_with_bindings q_constr_matching - q_hintdb; + q_hintdb q_move_location; anti: [ [ "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] ; @@ -692,6 +692,16 @@ GEXTEND Gram q_constr_matching: [ [ m = match_list -> m ] ] ; + move_location: + [ [ "at"; IDENT "top" -> Loc.tag ~loc:!@loc @@ QMoveFirst + | "at"; IDENT "bottom" -> Loc.tag ~loc:!@loc @@ QMoveLast + | IDENT "after"; id = ident_or_anti -> Loc.tag ~loc:!@loc @@ QMoveAfter id + | IDENT "before"; id = ident_or_anti -> Loc.tag ~loc:!@loc @@ QMoveBefore id + ] ] + ; + q_move_location: + [ [ mv = move_location -> mv ] ] + ; END (** Extension of constr syntax *) diff --git a/src/tac2core.ml b/src/tac2core.ml index a735dd19d9..37f1c99b15 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -1142,6 +1142,7 @@ let () = add_expr_scope "occurrences" q_occurrences Tac2quote.of_occurrences let () = add_expr_scope "dispatch" q_dispatch Tac2quote.of_dispatch let () = add_expr_scope "strategy" q_strategy_flag Tac2quote.of_strategy_flag let () = add_expr_scope "reference" q_reference Tac2quote.of_reference +let () = add_expr_scope "move_location" q_move_location Tac2quote.of_move_location let () = add_expr_scope "constr_matching" q_constr_matching Tac2quote.of_constr_matching let () = add_generic_scope "constr" Pcoq.Constr.constr Tac2quote.wit_constr diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 121841e8dc..9fd03ff5aa 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -39,6 +39,7 @@ let q_reference = Pcoq.Gram.entry_create "tactic:q_reference" let q_strategy_flag = Pcoq.Gram.entry_create "tactic:q_strategy_flag" let q_constr_matching = Pcoq.Gram.entry_create "tactic:q_constr_matching" let q_hintdb = Pcoq.Gram.entry_create "tactic:q_hintdb" +let q_move_location = Pcoq.Gram.entry_create "tactic:q_move_location" end (** Tactic definition *) diff --git a/src/tac2entries.mli b/src/tac2entries.mli index 91e2a683d8..dda1653593 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -74,6 +74,7 @@ val q_reference : Libnames.reference or_anti Pcoq.Gram.entry val q_strategy_flag : strategy_flag Pcoq.Gram.entry val q_constr_matching : constr_matching Pcoq.Gram.entry val q_hintdb : hintdb Pcoq.Gram.entry +val q_move_location : move_location Pcoq.Gram.entry end (** {5 Hooks} *) diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index 4bbaf43d8d..7d02022e07 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -131,3 +131,11 @@ type hintdb_r = | QHintDbs of Id.t located or_anti list type hintdb = hintdb_r located + +type move_location_r = +| QMoveAfter of Id.t located or_anti +| QMoveBefore of Id.t located or_anti +| QMoveFirst +| QMoveLast + +type move_location = move_location_r located diff --git a/src/tac2quote.ml b/src/tac2quote.ml index f87985435c..f14612d58f 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -364,3 +364,9 @@ let of_constr_matching (loc, m) = constructor ?loc (pattern_core "ConstrMatchContext") [pat; e] in of_list ?loc map m + +let of_move_location (loc, mv) = match mv with +| QMoveAfter id -> std_constructor ?loc "MoveAfter" [of_anti of_ident id] +| QMoveBefore id -> std_constructor ?loc "MoveBefore" [of_anti of_ident id] +| QMoveFirst -> std_constructor ?loc "MoveFirst" [] +| QMoveLast -> std_constructor ?loc "MoveLast" [] diff --git a/src/tac2quote.mli b/src/tac2quote.mli index b85f3438a3..db2fda3831 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -56,6 +56,8 @@ val of_occurrences : occurrences -> raw_tacexpr val of_hintdb : hintdb -> raw_tacexpr +val of_move_location : move_location -> raw_tacexpr + val of_reference : Libnames.reference or_anti -> raw_tacexpr val of_hyp : ?loc:Loc.t -> Id.t located -> raw_tacexpr diff --git a/theories/Std.v b/theories/Std.v index 02bc4ff450..3f98bdbaab 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -106,6 +106,13 @@ Ltac2 Type rewriting := { Ltac2 Type evar_flag := bool. Ltac2 Type advanced_flag := bool. +Ltac2 Type move_location := [ +| MoveAfter (ident) +| MoveBefore (ident) +| MoveFirst +| MoveLast +]. + Ltac2 Type inversion_kind := [ | SimpleInversion | FullInversion -- cgit v1.2.3 From c38e196fc175aaca2268f73107c9658c7af7d9fc Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 5 Sep 2017 15:59:42 +0200 Subject: Binding move and intro. --- src/tac2stdlib.ml | 20 ++++++++++++++++++++ theories/Notations.v | 5 +++++ theories/Std.v | 4 ++++ 3 files changed, 29 insertions(+) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 03141805ef..1762952e55 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -176,6 +176,13 @@ let to_inversion_kind = function | ValInt 2 -> Misctypes.FullInversionClear | _ -> assert false +let to_move_location = function +| ValInt 0 -> MoveFirst +| ValInt 1 -> MoveLast +| ValBlk (0, [|id|]) -> MoveAfter (Value.to_ident id) +| ValBlk (1, [|id|]) -> MoveBefore (Value.to_ident id) +| _ -> assert false + (** Standard tactics sharing their implementation with Ltac1 *) let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } @@ -488,6 +495,19 @@ end let () = define_prim0 "tac_reflexivity" Tactics.intros_reflexivity +let () = define_prim2 "tac_move" begin fun _ id mv -> + let id = Value.to_ident id in + let mv = to_move_location mv in + Tactics.move_hyp id mv +end + +let () = define_prim2 "tac_intro" begin fun _ id mv -> + let id = Value.to_option Value.to_ident id in + let mv = Value.to_option to_move_location mv in + let mv = Option.default MoveLast mv in + Tactics.intro_move id mv +end + (* TACTIC EXTEND exact diff --git a/theories/Notations.v b/theories/Notations.v index 8a3e769d12..9ecca018af 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -399,6 +399,11 @@ Ltac2 exact0 ev c := Ltac2 Notation "exact" c(thunk(open_constr)) := exact0 false c. Ltac2 Notation "eexact" c(thunk(open_constr)) := exact0 true c. +Ltac2 Notation "intro" id(opt(ident)) mv(opt(move_location)) := Std.intro id mv. +Ltac2 Notation intro := intro. + +Ltac2 Notation "move" id(ident) mv(move_location) := Std.move id mv. + Ltac2 Notation reflexivity := Std.reflexivity (). Ltac2 symmetry0 cl := diff --git a/theories/Std.v b/theories/Std.v index 3f98bdbaab..a937560b10 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -214,6 +214,10 @@ Ltac2 @ external inversion : inversion_kind -> destruction_arg -> intro_pattern (** coretactics *) +Ltac2 @ external move : ident -> move_location -> unit := "ltac2" "tac_move". + +Ltac2 @ external intro : ident option -> move_location option -> unit := "ltac2" "tac_intro". + (** extratactics *) Ltac2 @ external discriminate : evar_flag -> destruction_arg option -> unit := "ltac2" "tac_discriminate". -- cgit v1.2.3 From 68e803063818235acdc6ade35767ee618f88fe89 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 5 Sep 2017 16:27:44 +0200 Subject: The absurd tactic now parses a constr. --- theories/Notations.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/Notations.v b/theories/Notations.v index 9ecca018af..743210ae8d 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -435,7 +435,7 @@ Ltac2 Notation refine := Control.refine. Ltac2 absurd0 c := Control.enter (fun _ => Std.absurd (c ())). -Ltac2 Notation absurd := absurd0. +Ltac2 Notation "absurd" c(thunk(open_constr)) := absurd0 c. Ltac2 subst0 ids := match ids with | [] => Std.subst_all () -- cgit v1.2.3 From fb8142bb2fd84b10f4536fa63c972286365413f8 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 5 Sep 2017 16:31:16 +0200 Subject: Export Ltac2.Notations in the Ltac2 entry module. --- theories/Ltac2.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/Ltac2.v b/theories/Ltac2.v index 996236325c..7b2f592ac6 100644 --- a/theories/Ltac2.v +++ b/theories/Ltac2.v @@ -18,4 +18,4 @@ Require Ltac2.Control. Require Ltac2.Fresh. Require Ltac2.Pattern. Require Ltac2.Std. -Require Ltac2.Notations. +Require Export Ltac2.Notations. -- cgit v1.2.3 From 53c63e43a3daf99cf8bd44498b1c53798a8ba876 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 5 Sep 2017 17:19:41 +0200 Subject: Binding the firstorder tactic. --- src/tac2stdlib.ml | 9 +++++++++ src/tac2tactics.ml | 11 +++++++++++ src/tac2tactics.mli | 2 ++ theories/Notations.v | 22 +++++++++++++++++----- theories/Std.v | 4 ++++ 5 files changed, 43 insertions(+), 5 deletions(-) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 1762952e55..713a5f1b1c 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -705,3 +705,12 @@ let () = define_prim3 "tac_typeclasses_eauto" begin fun bt str n dbs -> let dbs = Value.to_option (fun l -> Value.to_list Value.to_ident l) dbs in Tac2tactics.typeclasses_eauto str n dbs end + +(** Firstorder *) + +let () = define_prim3 "tac_firstorder" begin fun bt tac refs ids -> + let tac = Value.to_option (fun t -> Proofview.tclIGNORE (thaw bt t)) tac in + let refs = Value.to_list Value.to_reference refs in + let ids = Value.to_list Value.to_ident ids in + Tac2tactics.firstorder tac refs ids +end diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index b069e57235..dacbb898d3 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -281,3 +281,14 @@ let inversion knd arg pat ids = end in on_destruction_arg inversion true (Some arg) + +(** Firstorder *) + +let firstorder tac refs ids = + let open Ground_plugin in + (** FUCK YOU API *) + let ids = List.map Id.to_string ids in + let tac : unit API.Proofview.tactic option = Obj.magic (tac : unit Proofview.tactic option) in + let refs : API.Globnames.global_reference list = Obj.magic (refs : Globnames.global_reference list) in + let ids : API.Hints.hint_db_name list = Obj.magic (ids : Hints.hint_db_name list) in + (Obj.magic (G_ground.gen_ground_tac true tac refs ids : unit API.Proofview.tactic) : unit Proofview.tactic) diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli index 8b466cd529..841e8fe762 100644 --- a/src/tac2tactics.mli +++ b/src/tac2tactics.mli @@ -101,3 +101,5 @@ val typeclasses_eauto : Class_tactics.search_strategy option -> int option -> Id.t list option -> unit Proofview.tactic val inversion : inversion_kind -> destruction_arg -> intro_pattern option -> Id.t list option -> unit tactic + +val firstorder : unit Proofview.tactic option -> global_reference list -> Id.t list -> unit tactic diff --git a/theories/Notations.v b/theories/Notations.v index 743210ae8d..9b39942ca5 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -470,14 +470,14 @@ Ltac2 default_db dbs := match dbs with end end. -Ltac2 default_using use := match use with +Ltac2 default_list use := match use with | None => [] | Some use => use end. Ltac2 trivial0 use dbs := let dbs := default_db dbs in - let use := default_using use in + let use := default_list use in Std.trivial Std.Off use dbs. Ltac2 Notation "trivial" @@ -488,7 +488,7 @@ Ltac2 Notation trivial := trivial. Ltac2 auto0 n use dbs := let dbs := default_db dbs in - let use := default_using use in + let use := default_list use in Std.auto Std.Off n use dbs. Ltac2 Notation "auto" n(opt(tactic(0))) @@ -499,7 +499,7 @@ Ltac2 Notation auto := auto. Ltac2 new_eauto0 n use dbs := let dbs := default_db dbs in - let use := default_using use in + let use := default_list use in Std.new_auto Std.Off n use dbs. Ltac2 Notation "new" "auto" n(opt(tactic(0))) @@ -508,7 +508,7 @@ Ltac2 Notation "new" "auto" n(opt(tactic(0))) Ltac2 eauto0 n p use dbs := let dbs := default_db dbs in - let use := default_using use in + let use := default_list use in Std.eauto Std.Off n p use dbs. Ltac2 Notation "eauto" n(opt(tactic(0))) p(opt(tactic(0))) @@ -529,3 +529,15 @@ Ltac2 Notation typeclasses_eauto := typeclasses_eauto. Ltac2 f_equal0 () := ltac1:(f_equal). Ltac2 Notation f_equal := f_equal0 (). + +(** Firstorder *) + +Ltac2 firstorder0 tac refs ids := + let refs := default_list refs in + let ids := default_list ids in + Std.firstorder tac refs ids. + +Ltac2 Notation "firstorder" + tac(opt(thunk(tactic))) + refs(opt(seq("using", list1(reference, ",")))) + ids(opt(seq("with", list1(ident)))) := firstorder0 tac refs ids. diff --git a/theories/Std.v b/theories/Std.v index a937560b10..f8b269dce6 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -246,3 +246,7 @@ Ltac2 @ external new_auto : debug -> int option -> (unit -> constr) list -> iden Ltac2 @ external eauto : debug -> int option -> int option -> (unit -> constr) list -> ident list option -> unit := "ltac2" "tac_eauto". Ltac2 @ external typeclasses_eauto : strategy option -> int option -> ident list option -> unit := "ltac2" "tac_typeclasses_eauto". + +(** firstorder *) + +Ltac2 @ external firstorder : (unit -> unit) option -> reference list -> ident list -> unit := "ltac2" "tac_firstorder". -- cgit v1.2.3 From 217bd80b651d2d12b05f74cf21485eb0fea8e3e3 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 5 Sep 2017 21:51:59 +0200 Subject: Refine does not evar-normalizes the goal preemptively. --- src/tac2core.ml | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 37f1c99b15..cd39cb6f27 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -734,7 +734,8 @@ end (** (unit -> constr) -> unit *) let () = define1 "refine" begin fun bt c -> let c = thaw bt c >>= fun c -> Proofview.tclUNIT ((), Value.to_constr c) in - Proofview.Goal.nf_enter begin fun gl -> + Proofview.Goal.enter begin fun gl -> + let gl = Proofview.Goal.assume gl in Refine.generic_refine ~typecheck:true c gl end >>= fun () -> return v_unit end -- cgit v1.2.3 From e4db6e726b5307c35e99a02cacbd96290e80dd24 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 6 Sep 2017 14:11:00 +0200 Subject: Code factorization. --- src/tac2core.ml | 7 +------ src/tac2quote.mli | 2 ++ 2 files changed, 3 insertions(+), 6 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index cd39cb6f27..2c9226adae 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -1012,11 +1012,6 @@ let scope_fail s args = let q_unit = Loc.tag @@ CTacCst (AbsKn (Tuple 0)) -let rthunk e = - let loc = Tac2intern.loc_of_tacexpr e in - let var = [Loc.tag ?loc @@ CPatVar Anonymous, Some (Loc.tag ?loc @@ CTypRef (AbsKn (Other Core.t_unit), []))] in - Loc.tag ?loc @@ CTacFun (var, e) - let add_generic_scope s entry arg = let parse = function | [] -> @@ -1118,7 +1113,7 @@ end let () = add_scope "thunk" begin function | [tok] -> let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in - let act e = rthunk (act e) in + let act e = Tac2quote.thunk (act e) in Tac2entries.ScopeRule (scope, act) | arg -> scope_fail "thunk" arg end diff --git a/src/tac2quote.mli b/src/tac2quote.mli index db2fda3831..9f42c60042 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -20,6 +20,8 @@ open Tac2expr val constructor : ?loc:Loc.t -> ltac_constructor -> raw_tacexpr list -> raw_tacexpr +val thunk : raw_tacexpr -> raw_tacexpr + val of_anti : ('a -> raw_tacexpr) -> 'a or_anti -> raw_tacexpr val of_int : int located -> raw_tacexpr -- cgit v1.2.3 From a36ec5034231b2f879538bcb5c8401d03f2ad04f Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 6 Sep 2017 18:23:54 +0200 Subject: Introducing abstract data representations. --- src/tac2ffi.ml | 76 +++++++++++++++++++++++++++++++++++++++++++++++++++++++++ src/tac2ffi.mli | 23 +++++++++++++++++ 2 files changed, 99 insertions(+) diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index 4ed0096787..6a60562d1a 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -12,6 +12,12 @@ open Genarg open Tac2dyn open Tac2expr +type 'a repr = { + r_of : 'a -> valexpr; + r_to : valexpr -> 'a; + r_id : bool; +} + (** Dynamic tags *) let val_exn = Val.create "exn" @@ -44,11 +50,23 @@ let to_unit = function | ValInt 0 -> () | _ -> assert false +let unit = { + r_of = of_unit; + r_to = to_unit; + r_id = false; +} + let of_int n = ValInt n let to_int = function | ValInt n -> n | _ -> assert false +let int = { + r_of = of_int; + r_to = to_int; + r_id = false; +} + let of_bool b = if b then ValInt 0 else ValInt 1 let to_bool = function @@ -56,16 +74,34 @@ let to_bool = function | ValInt 1 -> false | _ -> assert false +let bool = { + r_of = of_bool; + r_to = to_bool; + r_id = false; +} + let of_char n = ValInt (Char.code n) let to_char = function | ValInt n -> Char.chr n | _ -> assert false +let char = { + r_of = of_char; + r_to = to_char; + r_id = false; +} + let of_string s = ValStr s let to_string = function | ValStr s -> s | _ -> assert false +let string = { + r_of = of_string; + r_to = to_string; + r_id = false; +} + let rec of_list f = function | [] -> ValInt 0 | x :: l -> ValBlk (0, [| f x; of_list f l |]) @@ -75,6 +111,12 @@ let rec to_list f = function | ValBlk (0, [|v; vl|]) -> f v :: to_list f vl | _ -> assert false +let list r = { + r_of = (fun l -> of_list r.r_of l); + r_to = (fun l -> to_list r.r_to l); + r_id = false; +} + let of_ext tag c = ValExt (tag, c) @@ -82,14 +124,23 @@ let to_ext tag = function | ValExt (tag', e) -> extract_val tag tag' e | _ -> assert false +let repr_ext tag = { + r_of = (fun e -> of_ext tag e); + r_to = (fun e -> to_ext tag e); + r_id = false; +} + let of_constr c = of_ext val_constr c let to_constr c = to_ext val_constr c +let constr = repr_ext val_constr let of_ident c = of_ext val_ident c let to_ident c = to_ext val_ident c +let ident = repr_ext val_ident let of_pattern c = of_ext val_pattern c let to_pattern c = to_ext val_pattern c +let pattern = repr_ext val_pattern let internal_err = let open Names in @@ -108,6 +159,12 @@ let to_exn c = match c with (Tac2interp.LtacError (kn, c, []), Exninfo.null) | _ -> assert false +let exn = { + r_of = of_exn; + r_to = to_exn; + r_id = false; +} + let of_option f = function | None -> ValInt 0 | Some c -> ValBlk (0, [|f c|]) @@ -117,8 +174,15 @@ let to_option f = function | ValBlk (0, [|c|]) -> Some (f c) | _ -> assert false +let option r = { + r_of = (fun l -> of_option r.r_of l); + r_to = (fun l -> to_option r.r_to l); + r_id = false; +} + let of_pp c = of_ext val_pp c let to_pp c = to_ext val_pp c +let pp = repr_ext val_pp let of_tuple cl = ValBlk (0, cl) let to_tuple = function @@ -129,9 +193,15 @@ let of_array f vl = ValBlk (0, Array.map f vl) let to_array f = function | ValBlk (0, vl) -> Array.map f vl | _ -> assert false +let array r = { + r_of = (fun l -> of_array r.r_of l); + r_to = (fun l -> to_array r.r_to l); + r_id = false; +} let of_constant c = of_ext val_constant c let to_constant c = to_ext val_constant c +let constant = repr_ext val_constant let of_reference = function | VarRef id -> ValBlk (0, [| of_ident id |]) @@ -145,3 +215,9 @@ let to_reference = function | ValBlk (2, [| ind |]) -> IndRef (to_ext val_inductive ind) | ValBlk (3, [| cstr |]) -> ConstructRef (to_ext val_constructor cstr) | _ -> assert false + +let reference = { + r_of = of_reference; + r_to = to_reference; + r_id = false; +} diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index e836319349..db3087534b 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -13,60 +13,83 @@ open Tac2expr (** {5 Ltac2 FFI} *) +type 'a repr = { + r_of : 'a -> valexpr; + r_to : valexpr -> 'a; + r_id : bool; + (** True if the functions above are physical identities. *) +} + (** These functions allow to convert back and forth between OCaml and Ltac2 data representation. The [to_*] functions raise an anomaly whenever the data has not expected shape. *) val of_unit : unit -> valexpr val to_unit : valexpr -> unit +val unit : unit repr val of_int : int -> valexpr val to_int : valexpr -> int +val int : int repr val of_bool : bool -> valexpr val to_bool : valexpr -> bool +val bool : bool repr val of_char : char -> valexpr val to_char : valexpr -> char +val char : char repr val of_string : string -> valexpr val to_string : valexpr -> string +val string : string repr val of_list : ('a -> valexpr) -> 'a list -> valexpr val to_list : (valexpr -> 'a) -> valexpr -> 'a list +val list : 'a repr -> 'a list repr val of_constr : EConstr.t -> valexpr val to_constr : valexpr -> EConstr.t +val constr : EConstr.t repr val of_exn : Exninfo.iexn -> valexpr val to_exn : valexpr -> Exninfo.iexn +val exn : Exninfo.iexn repr val of_ident : Id.t -> valexpr val to_ident : valexpr -> Id.t +val ident : Id.t repr val of_array : ('a -> valexpr) -> 'a array -> valexpr val to_array : (valexpr -> 'a) -> valexpr -> 'a array +val array : 'a repr -> 'a array repr val of_tuple : valexpr array -> valexpr val to_tuple : valexpr -> valexpr array val of_option : ('a -> valexpr) -> 'a option -> valexpr val to_option : (valexpr -> 'a) -> valexpr -> 'a option +val option : 'a repr -> 'a option repr val of_pattern : Pattern.constr_pattern -> valexpr val to_pattern : valexpr -> Pattern.constr_pattern +val pattern : Pattern.constr_pattern repr val of_pp : Pp.t -> valexpr val to_pp : valexpr -> Pp.t +val pp : Pp.t repr val of_constant : Constant.t -> valexpr val to_constant : valexpr -> Constant.t +val constant : Constant.t repr val of_reference : Globnames.global_reference -> valexpr val to_reference : valexpr -> Globnames.global_reference +val reference : Globnames.global_reference repr val of_ext : 'a Val.tag -> 'a -> valexpr val to_ext : 'a Val.tag -> valexpr -> 'a +val repr_ext : 'a Val.tag -> 'a repr (** {5 Dynamic tags} *) -- cgit v1.2.3 From 6985a2001d28ff0850198d8219d7b791a226bdac Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 6 Sep 2017 18:53:07 +0200 Subject: Moving Tac2ffi before Tac2interp. --- src/ltac2_plugin.mlpack | 2 +- src/tac2ffi.ml | 9 +++++++-- src/tac2ffi.mli | 5 +++++ src/tac2interp.ml | 2 +- 4 files changed, 14 insertions(+), 4 deletions(-) diff --git a/src/ltac2_plugin.mlpack b/src/ltac2_plugin.mlpack index 92f391a085..00ba5bc58e 100644 --- a/src/ltac2_plugin.mlpack +++ b/src/ltac2_plugin.mlpack @@ -2,9 +2,9 @@ Tac2dyn Tac2env Tac2print Tac2intern +Tac2ffi Tac2interp Tac2entries -Tac2ffi Tac2quote Tac2core Tac2tactics diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index 6a60562d1a..a0c51783ed 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -7,6 +7,7 @@ (************************************************************************) open Util +open Names open Globnames open Genarg open Tac2dyn @@ -42,6 +43,10 @@ match Val.eq tag tag' with | None -> assert false | Some Refl -> v +(** Exception *) + +exception LtacError of KerName.t * valexpr array * backtrace + (** Conversion functions *) let of_unit () = ValInt 0 @@ -148,7 +153,7 @@ let internal_err = (** FIXME: handle backtrace in Ltac2 exceptions *) let of_exn c = match fst c with -| Tac2interp.LtacError (kn, c, _) -> ValOpn (kn, c) +| LtacError (kn, c, _) -> ValOpn (kn, c) | _ -> ValOpn (internal_err, [|of_ext val_exn c|]) let to_exn c = match c with @@ -156,7 +161,7 @@ let to_exn c = match c with if Names.KerName.equal kn internal_err then to_ext val_exn c.(0) else - (Tac2interp.LtacError (kn, c, []), Exninfo.null) + (LtacError (kn, c, []), Exninfo.null) | _ -> assert false let exn = { diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index db3087534b..2282d34b3a 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -111,3 +111,8 @@ val val_free : Id.Set.t Val.tag val val_exn : Exninfo.iexn Tac2dyn.Val.tag (** Toplevel representation of OCaml exceptions. Invariant: no [LtacError] should be put into a value with tag [val_exn]. *) + +(** Exception *) + +exception LtacError of KerName.t * valexpr array * backtrace +(** Ltac2-defined exceptions seen from OCaml side *) diff --git a/src/tac2interp.ml b/src/tac2interp.ml index b58ce6b851..e79ce87268 100644 --- a/src/tac2interp.ml +++ b/src/tac2interp.ml @@ -14,7 +14,7 @@ open Names open Proofview.Notations open Tac2expr -exception LtacError of KerName.t * valexpr array * backtrace +exception LtacError = Tac2ffi.LtacError let empty_environment = { env_ist = Id.Map.empty; -- cgit v1.2.3 From 841c4a028b5cf7e3cfff6b91a33db38a4b8d54df Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 6 Sep 2017 18:55:27 +0200 Subject: The interp_app function now takes a closure as an argument. --- src/tac2core.ml | 5 +++-- src/tac2ffi.ml | 12 ++++++++++++ src/tac2ffi.mli | 4 ++++ src/tac2interp.ml | 9 ++++----- src/tac2interp.mli | 2 +- src/tac2stdlib.ml | 2 +- 6 files changed, 25 insertions(+), 9 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 2c9226adae..5e475e1d4a 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -103,7 +103,7 @@ let err_matchfailure bt = (** Helper functions *) -let thaw bt f = Tac2interp.interp_app bt f [v_unit] +let thaw bt f = Tac2interp.interp_app bt (Value.to_closure f) [v_unit] let throw bt e = Proofview.tclLIFT (Proofview.NonLogical.raise (e bt)) let set_bt bt e = match e with @@ -615,7 +615,7 @@ end (** (unit -> 'a) -> (exn -> 'a) -> 'a *) let () = define2 "plus" begin fun bt x k -> - Proofview.tclOR (thaw bt x) (fun e -> Tac2interp.interp_app bt k [Value.of_exn e]) + Proofview.tclOR (thaw bt x) (fun e -> Tac2interp.interp_app bt (Value.to_closure k) [Value.of_exn e]) end (** (unit -> 'a) -> 'a *) @@ -741,6 +741,7 @@ let () = define1 "refine" begin fun bt c -> end let () = define2 "with_holes" begin fun bt x f -> + let f = Value.to_closure f in Proofview.tclEVARMAP >>= fun sigma0 -> thaw bt x >>= fun ans -> Proofview.tclEVARMAP >>= fun sigma -> diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index a0c51783ed..c3f535c1bc 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -122,6 +122,18 @@ let list r = { r_id = false; } +let of_closure cls = ValCls cls + +let to_closure = function +| ValCls cls -> cls +| ValExt _ | ValInt _ | ValBlk _ | ValStr _ | ValOpn _ -> assert false + +let closure = { + r_of = of_closure; + r_to = to_closure; + r_id = false; +} + let of_ext tag c = ValExt (tag, c) diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index 2282d34b3a..fe813b0e35 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -60,6 +60,10 @@ val of_ident : Id.t -> valexpr val to_ident : valexpr -> Id.t val ident : Id.t repr +val of_closure : closure -> valexpr +val to_closure : valexpr -> closure +val closure : closure repr + val of_array : ('a -> valexpr) -> 'a array -> valexpr val to_array : (valexpr -> 'a) -> valexpr -> 'a array val array : 'a repr -> 'a array repr diff --git a/src/tac2interp.ml b/src/tac2interp.ml index e79ce87268..6bee5a0794 100644 --- a/src/tac2interp.ml +++ b/src/tac2interp.ml @@ -46,7 +46,7 @@ let rec interp (ist : environment) = function | GTacApp (f, args) -> interp ist f >>= fun f -> Proofview.Monad.List.map (fun e -> interp ist e) args >>= fun args -> - interp_app ist.env_bkt f args + interp_app ist.env_bkt (Tac2ffi.to_closure f) args | GTacLet (false, el, e) -> let fold accu (na, e) = interp ist e >>= fun e -> @@ -97,18 +97,17 @@ let rec interp (ist : environment) = function tpe.Tac2env.ml_interp ist e and interp_app bt f args = match f with -| ValCls { clos_env = ist; clos_var = ids; clos_exp = e; clos_ref = kn } -> +| { clos_env = ist; clos_var = ids; clos_exp = e; clos_ref = kn } -> let rec push ist ids args = match ids, args with | [], [] -> interp ist e - | [], _ :: _ -> interp ist e >>= fun f -> interp_app bt f args + | [], _ :: _ -> + interp ist e >>= fun f -> interp_app bt (Tac2ffi.to_closure f) args | _ :: _, [] -> let cls = { clos_ref = kn; clos_env = ist.env_ist; clos_var = ids; clos_exp = e } in return (ValCls cls) | id :: ids, arg :: args -> push (push_name ist id arg) ids args in push { env_ist = ist; env_bkt = FrLtac kn :: bt } ids args -| ValExt _ | ValInt _ | ValBlk _ | ValStr _ | ValOpn _ -> - anomaly (str "Unexpected value shape") and interp_case ist e cse0 cse1 = match e with | ValInt n -> interp ist cse0.(n) diff --git a/src/tac2interp.mli b/src/tac2interp.mli index ea7db33b60..ca263b2c4b 100644 --- a/src/tac2interp.mli +++ b/src/tac2interp.mli @@ -13,7 +13,7 @@ val empty_environment : environment val interp : environment -> glb_tacexpr -> valexpr Proofview.tactic -val interp_app : backtrace -> valexpr -> valexpr list -> valexpr Proofview.tactic +val interp_app : backtrace -> closure -> valexpr list -> valexpr Proofview.tactic (** {5 Cross-boundary encodings} *) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 713a5f1b1c..79af41b7d0 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -18,7 +18,7 @@ module Value = Tac2ffi let return x = Proofview.tclUNIT x let v_unit = Value.of_unit () -let thaw bt f = Tac2interp.interp_app bt f [v_unit] +let thaw bt f = Tac2interp.interp_app bt (Value.to_closure f) [v_unit] let to_pair f g = function | ValBlk (0, [| x; y |]) -> (f x, g y) -- cgit v1.2.3 From f5ed96350ecc947ad4e55be9439cd0d30c68bde0 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 6 Sep 2017 19:03:32 +0200 Subject: Parameterizing over parameters in ML functions from Tac2core. --- src/tac2core.ml | 213 ++++++++++++++++++++------------------------------------ src/tac2ffi.ml | 12 ++++ src/tac2ffi.mli | 4 ++ 3 files changed, 93 insertions(+), 136 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 5e475e1d4a..2f349e32af 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -20,6 +20,7 @@ open Proofview.Notations (** Standard values *) module Value = Tac2ffi +open Value let std_core n = KerName.make2 Tac2env.std_prefix (Label.of_id (Id.of_string_soft n)) let coq_core n = KerName.make2 Tac2env.coq_prefix (Label.of_id (Id.of_string_soft n)) @@ -53,10 +54,6 @@ open Core let v_unit = ValInt 0 -let to_block = function -| ValBlk (_, v) -> v -| _ -> assert false - let of_name c = match c with | Anonymous -> Value.of_option Value.of_ident None | Name id -> Value.of_option Value.of_ident (Some id) @@ -103,7 +100,8 @@ let err_matchfailure bt = (** Helper functions *) -let thaw bt f = Tac2interp.interp_app bt (Value.to_closure f) [v_unit] +let thaw bt f = Tac2interp.interp_app bt f [v_unit] + let throw bt e = Proofview.tclLIFT (Proofview.NonLogical.raise (e bt)) let set_bt bt e = match e with @@ -144,111 +142,95 @@ let define0 name f = Tac2env.define_primitive (pname name) begin fun bt arg -> m | _ -> assert false end -let define1 name f = Tac2env.define_primitive (pname name) begin fun bt arg -> match arg with -| [x] -> f bt x +let define1 name r0 f = Tac2env.define_primitive (pname name) begin fun bt arg -> match arg with +| [x] -> f bt (r0.Value.r_to x) | _ -> assert false end -let define2 name f = Tac2env.define_primitive (pname name) begin fun bt arg -> match arg with -| [x; y] -> f bt x y +let define2 name r0 r1 f = Tac2env.define_primitive (pname name) begin fun bt arg -> match arg with +| [x; y] -> f bt (r0.Value.r_to x) (r1.Value.r_to y) | _ -> assert false end -let define3 name f = Tac2env.define_primitive (pname name) begin fun bt arg -> match arg with -| [x; y; z] -> f bt x y z +let define3 name r0 r1 r2 f = Tac2env.define_primitive (pname name) begin fun bt arg -> match arg with +| [x; y; z] -> f bt (r0.Value.r_to x) (r1.Value.r_to y) (r2.Value.r_to z) | _ -> assert false end (** Printing *) -let () = define1 "print" begin fun _ pp -> - wrap_unit (fun () -> Feedback.msg_notice (Value.to_pp pp)) +let () = define1 "print" pp begin fun _ pp -> + wrap_unit (fun () -> Feedback.msg_notice pp) end -let () = define1 "message_of_int" begin fun _ n -> - let n = Value.to_int n in - return (Value.of_pp (int n)) +let () = define1 "message_of_int" int begin fun _ n -> + return (Value.of_pp (Pp.int n)) end -let () = define1 "message_of_string" begin fun _ s -> - let s = Value.to_string s in +let () = define1 "message_of_string" string begin fun _ s -> return (Value.of_pp (str (Bytes.to_string s))) end -let () = define1 "message_of_constr" begin fun bt c -> +let () = define1 "message_of_constr" constr begin fun bt c -> pf_apply bt begin fun env sigma -> - let c = Value.to_constr c in let pp = Printer.pr_econstr_env env sigma c in return (Value.of_pp pp) end end -let () = define1 "message_of_ident" begin fun _ c -> - let c = Value.to_ident c in +let () = define1 "message_of_ident" ident begin fun _ c -> let pp = Id.print c in return (Value.of_pp pp) end -let () = define2 "message_concat" begin fun _ m1 m2 -> - let m1 = Value.to_pp m1 in - let m2 = Value.to_pp m2 in +let () = define2 "message_concat" pp pp begin fun _ m1 m2 -> return (Value.of_pp (Pp.app m1 m2)) end (** Array *) -let () = define2 "array_make" begin fun bt n x -> - let n = Value.to_int n in +let () = define2 "array_make" int valexpr begin fun bt n x -> if n < 0 || n > Sys.max_array_length then throw bt err_outofbounds else wrap (fun () -> ValBlk (0, Array.make n x)) end -let () = define1 "array_length" begin fun _ v -> - let v = to_block v in +let () = define1 "array_length" block begin fun _ v -> return (ValInt (Array.length v)) end -let () = define3 "array_set" begin fun bt v n x -> - let v = to_block v in - let n = Value.to_int n in +let () = define3 "array_set" block int valexpr begin fun bt v n x -> if n < 0 || n >= Array.length v then throw bt err_outofbounds else wrap_unit (fun () -> v.(n) <- x) end -let () = define2 "array_get" begin fun bt v n -> - let v = to_block v in - let n = Value.to_int n in +let () = define2 "array_get" block int begin fun bt v n -> if n < 0 || n >= Array.length v then throw bt err_outofbounds else wrap (fun () -> v.(n)) end (** Ident *) -let () = define2 "ident_equal" begin fun _ id1 id2 -> - let id1 = Value.to_ident id1 in - let id2 = Value.to_ident id2 in +let () = define2 "ident_equal" ident ident begin fun _ id1 id2 -> return (Value.of_bool (Id.equal id1 id2)) end -let () = define1 "ident_to_string" begin fun _ id -> - let id = Value.to_ident id in +let () = define1 "ident_to_string" ident begin fun _ id -> return (Value.of_string (Id.to_string id)) end -let () = define1 "ident_of_string" begin fun _ s -> - let s = Value.to_string s in +let () = define1 "ident_of_string" string begin fun _ s -> let id = try Some (Id.of_string s) with _ -> None in return (Value.of_option Value.of_ident id) end (** Int *) -let () = define2 "int_equal" begin fun _ m n -> - return (Value.of_bool (Value.to_int m == Value.to_int n)) +let () = define2 "int_equal" int int begin fun _ m n -> + return (Value.of_bool (m == n)) end -let binop n f = define2 n begin fun _ m n -> - return (Value.of_int (f (Value.to_int m) (Value.to_int n))) +let binop n f = define2 n int int begin fun _ m n -> + return (Value.of_int (f m m)) end let () = binop "int_compare" Int.compare @@ -256,34 +238,27 @@ let () = binop "int_add" (+) let () = binop "int_sub" (-) let () = binop "int_mul" ( * ) -let () = define1 "int_neg" begin fun _ m -> - return (Value.of_int (~- (Value.to_int m))) +let () = define1 "int_neg" int begin fun _ m -> + return (Value.of_int (~- m)) end (** String *) -let () = define2 "string_make" begin fun bt n c -> - let n = Value.to_int n in - let c = Value.to_char c in +let () = define2 "string_make" int char begin fun bt n c -> if n < 0 || n > Sys.max_string_length then throw bt err_outofbounds else wrap (fun () -> Value.of_string (Bytes.make n c)) end -let () = define1 "string_length" begin fun _ s -> - return (Value.of_int (Bytes.length (Value.to_string s))) +let () = define1 "string_length" string begin fun _ s -> + return (Value.of_int (Bytes.length s)) end -let () = define3 "string_set" begin fun bt s n c -> - let s = Value.to_string s in - let n = Value.to_int n in - let c = Value.to_char c in +let () = define3 "string_set" string int char begin fun bt s n c -> if n < 0 || n >= Bytes.length s then throw bt err_outofbounds else wrap_unit (fun () -> Bytes.set s n c) end -let () = define2 "string_get" begin fun bt s n -> - let s = Value.to_string s in - let n = Value.to_int n in +let () = define2 "string_get" string int begin fun bt s n -> if n < 0 || n >= Bytes.length s then throw bt err_outofbounds else wrap (fun () -> Value.of_char (Bytes.get s n)) end @@ -291,8 +266,7 @@ end (** Terms *) (** constr -> constr *) -let () = define1 "constr_type" begin fun bt c -> - let c = Value.to_constr c in +let () = define1 "constr_type" constr begin fun bt c -> let get_type env sigma = Proofview.V82.wrap_exceptions begin fun () -> let (sigma, t) = Typing.type_of env sigma c in @@ -303,18 +277,15 @@ let () = define1 "constr_type" begin fun bt c -> end (** constr -> constr *) -let () = define2 "constr_equal" begin fun _ c1 c2 -> - let c1 = Value.to_constr c1 in - let c2 = Value.to_constr c2 in +let () = define2 "constr_equal" constr constr begin fun _ c1 c2 -> Proofview.tclEVARMAP >>= fun sigma -> let b = EConstr.eq_constr sigma c1 c2 in Proofview.tclUNIT (Value.of_bool b) end -let () = define1 "constr_kind" begin fun _ c -> +let () = define1 "constr_kind" constr begin fun _ c -> let open Constr in Proofview.tclEVARMAP >>= fun sigma -> - let c = Value.to_constr c in return begin match EConstr.kind sigma c with | Rel n -> ValBlk (0, [|Value.of_int n|]) @@ -406,7 +377,7 @@ let () = define1 "constr_kind" begin fun _ c -> end end -let () = define1 "constr_make" begin fun _ knd -> +let () = define1 "constr_make" valexpr begin fun _ knd -> let open Constr in let c = match knd with | ValBlk (0, [|n|]) -> @@ -486,8 +457,7 @@ let () = define1 "constr_make" begin fun _ knd -> return (Value.of_constr c) end -let () = define1 "constr_check" begin fun bt c -> - let c = Value.to_constr c in +let () = define1 "constr_check" constr begin fun bt c -> pf_apply bt begin fun env sigma -> try let (sigma, _) = Typing.type_of env sigma c in @@ -499,27 +469,19 @@ let () = define1 "constr_check" begin fun bt c -> end end -let () = define3 "constr_substnl" begin fun _ subst k c -> - let subst = Value.to_list Value.to_constr subst in - let k = Value.to_int k in - let c = Value.to_constr c in +let () = define3 "constr_substnl" (list constr) int constr begin fun _ subst k c -> let ans = EConstr.Vars.substnl subst k c in return (Value.of_constr ans) end -let () = define3 "constr_closenl" begin fun _ ids k c -> - let ids = Value.to_list Value.to_ident ids in - let k = Value.to_int k in - let c = Value.to_constr c in +let () = define3 "constr_closenl" (list ident) int constr begin fun _ ids k c -> let ans = EConstr.Vars.substn_vars k ids c in return (Value.of_constr ans) end (** Patterns *) -let () = define2 "pattern_matches" begin fun bt pat c -> - let pat = Value.to_pattern pat in - let c = Value.to_constr c in +let () = define2 "pattern_matches" pattern constr begin fun bt pat c -> pf_apply bt begin fun env sigma -> let ans = try Some (Constr_matching.matches env sigma pat c) @@ -535,9 +497,7 @@ let () = define2 "pattern_matches" begin fun bt pat c -> end end -let () = define2 "pattern_matches_subterm" begin fun bt pat c -> - let pat = Value.to_pattern pat in - let c = Value.to_constr c in +let () = define2 "pattern_matches_subterm" pattern constr begin fun bt pat c -> let open Constr_matching in let rec of_ans s = match IStream.peek s with | IStream.Nil -> Proofview.tclZERO (err_matchfailure bt) @@ -553,9 +513,7 @@ let () = define2 "pattern_matches_subterm" begin fun bt pat c -> end end -let () = define2 "pattern_matches_vect" begin fun bt pat c -> - let pat = Value.to_pattern pat in - let c = Value.to_constr c in +let () = define2 "pattern_matches_vect" pattern constr begin fun bt pat c -> pf_apply bt begin fun env sigma -> let ans = try Some (Constr_matching.matches env sigma pat c) @@ -571,9 +529,7 @@ let () = define2 "pattern_matches_vect" begin fun bt pat c -> end end -let () = define2 "pattern_matches_subterm_vect" begin fun bt pat c -> - let pat = Value.to_pattern pat in - let c = Value.to_constr c in +let () = define2 "pattern_matches_subterm_vect" pattern constr begin fun bt pat c -> let open Constr_matching in let rec of_ans s = match IStream.peek s with | IStream.Nil -> Proofview.tclZERO (err_matchfailure bt) @@ -589,17 +545,16 @@ let () = define2 "pattern_matches_subterm_vect" begin fun bt pat c -> end end -let () = define2 "pattern_instantiate" begin fun _ ctx c -> - let ctx = EConstr.Unsafe.to_constr (Value.to_constr ctx) in - let c = EConstr.Unsafe.to_constr (Value.to_constr c) in +let () = define2 "pattern_instantiate" constr constr begin fun _ ctx c -> + let ctx = EConstr.Unsafe.to_constr ctx in + let c = EConstr.Unsafe.to_constr c in let ans = Termops.subst_meta [Constr_matching.special_meta, c] ctx in return (Value.of_constr (EConstr.of_constr ans)) end (** Error *) -let () = define1 "throw" begin fun bt e -> - let (e, info) = Value.to_exn e in +let () = define1 "throw" exn begin fun bt (e, info) -> let e = set_bt bt e in Proofview.tclLIFT (Proofview.NonLogical.raise ~info e) end @@ -607,38 +562,37 @@ end (** Control *) (** exn -> 'a *) -let () = define1 "zero" begin fun bt e -> - let (e, info) = Value.to_exn e in +let () = define1 "zero" exn begin fun bt (e, info) -> let e = set_bt bt e in Proofview.tclZERO ~info e end (** (unit -> 'a) -> (exn -> 'a) -> 'a *) -let () = define2 "plus" begin fun bt x k -> - Proofview.tclOR (thaw bt x) (fun e -> Tac2interp.interp_app bt (Value.to_closure k) [Value.of_exn e]) +let () = define2 "plus" closure closure begin fun bt x k -> + Proofview.tclOR (thaw bt x) (fun e -> Tac2interp.interp_app bt k [Value.of_exn e]) end (** (unit -> 'a) -> 'a *) -let () = define1 "once" begin fun bt f -> +let () = define1 "once" closure begin fun bt f -> Proofview.tclONCE (thaw bt f) end (** (unit -> unit) list -> unit *) -let () = define1 "dispatch" begin fun bt l -> - let l = Value.to_list (fun f -> Proofview.tclIGNORE (thaw bt f)) l in +let () = define1 "dispatch" (list closure) begin fun bt l -> + let l = List.map (fun f -> Proofview.tclIGNORE (thaw bt f)) l in Proofview.tclDISPATCH l >>= fun () -> return v_unit end (** (unit -> unit) list -> (unit -> unit) -> (unit -> unit) list -> unit *) -let () = define3 "extend" begin fun bt lft tac rgt -> - let lft = Value.to_list (fun f -> Proofview.tclIGNORE (thaw bt f)) lft in +let () = define3 "extend" (list closure) closure (list closure) begin fun bt lft tac rgt -> + let lft = List.map (fun f -> Proofview.tclIGNORE (thaw bt f)) lft in let tac = Proofview.tclIGNORE (thaw bt tac) in - let rgt = Value.to_list (fun f -> Proofview.tclIGNORE (thaw bt f)) rgt in + let rgt = List.map (fun f -> Proofview.tclIGNORE (thaw bt f)) rgt in Proofview.tclEXTEND lft tac rgt >>= fun () -> return v_unit end (** (unit -> unit) -> unit *) -let () = define1 "enter" begin fun bt f -> +let () = define1 "enter" closure begin fun bt f -> let f = Proofview.tclIGNORE (thaw bt f) in Proofview.tclINDEPENDENT f >>= fun () -> return v_unit end @@ -648,7 +602,7 @@ let e_var = Id.of_string "e" let prm_apply_kont_h = pname "apply_kont" (** (unit -> 'a) -> ('a * ('exn -> 'a)) result *) -let () = define1 "case" begin fun bt f -> +let () = define1 "case" closure begin fun bt f -> Proofview.tclCASE (thaw bt f) >>= begin function | Proofview.Next (x, k) -> let k = { @@ -663,16 +617,13 @@ let () = define1 "case" begin fun bt f -> end (** 'a kont -> exn -> 'a *) -let () = define2 "apply_kont" begin fun bt k e -> - let (e, info) = Value.to_exn e in +let () = define2 "apply_kont" (repr_ext val_kont) exn begin fun bt k (e, info) -> let e = set_bt bt e in - (Value.to_ext Value.val_kont k) (e, info) + k (e, info) end (** int -> int -> (unit -> 'a) -> 'a *) -let () = define3 "focus" begin fun bt i j tac -> - let i = Value.to_int i in - let j = Value.to_int j in +let () = define3 "focus" int int closure begin fun bt i j tac -> Proofview.tclFOCUS i j (thaw bt tac) end @@ -686,8 +637,8 @@ let () = define0 "shelve_unifiable" begin fun _ -> Proofview.shelve_unifiable >>= fun () -> return v_unit end -let () = define1 "new_goal" begin fun bt ev -> - let ev = Evar.unsafe_of_int (Value.to_int ev) in +let () = define1 "new_goal" int begin fun bt ev -> + let ev = Evar.unsafe_of_int ev in Proofview.tclEVARMAP >>= fun sigma -> if Evd.mem sigma ev then Proofview.Unsafe.tclNEWGOALS [ev] <*> Proofview.tclUNIT v_unit @@ -704,8 +655,7 @@ let () = define0 "goal" begin fun bt -> end (** ident -> constr *) -let () = define1 "hyp" begin fun bt id -> - let id = Value.to_ident id in +let () = define1 "hyp" ident begin fun bt id -> pf_apply bt begin fun env _ -> let mem = try ignore (Environ.lookup_named id env); true with Not_found -> false in if mem then return (Value.of_constr (EConstr.mkVar id)) @@ -732,7 +682,7 @@ let () = define0 "hyps" begin fun bt -> end (** (unit -> constr) -> unit *) -let () = define1 "refine" begin fun bt c -> +let () = define1 "refine" closure begin fun bt c -> let c = thaw bt c >>= fun c -> Proofview.tclUNIT ((), Value.to_constr c) in Proofview.Goal.enter begin fun gl -> let gl = Proofview.Goal.assume gl in @@ -740,8 +690,7 @@ let () = define1 "refine" begin fun bt c -> end >>= fun () -> return v_unit end -let () = define2 "with_holes" begin fun bt x f -> - let f = Value.to_closure f in +let () = define2 "with_holes" closure closure begin fun bt x f -> Proofview.tclEVARMAP >>= fun sigma0 -> thaw bt x >>= fun ans -> Proofview.tclEVARMAP >>= fun sigma -> @@ -749,18 +698,16 @@ let () = define2 "with_holes" begin fun bt x f -> Tacticals.New.tclWITHHOLES false (Tac2interp.interp_app bt f [ans]) sigma end -let () = define1 "progress" begin fun bt f -> +let () = define1 "progress" closure begin fun bt f -> Proofview.tclPROGRESS (thaw bt f) end -let () = define2 "abstract" begin fun bt id f -> - let id = Value.to_option Value.to_ident id in +let () = define2 "abstract" (option ident) closure begin fun bt id f -> Tactics.tclABSTRACT id (Proofview.tclIGNORE (thaw bt f)) >>= fun () -> return v_unit end -let () = define2 "time" begin fun bt s f -> - let s = Value.to_option Value.to_string s in +let () = define2 "time" (option string) closure begin fun bt s f -> Proofview.tclTIME s (thaw bt f) end @@ -770,21 +717,17 @@ end (** Fresh *) -let () = define2 "fresh_free_union" begin fun _ set1 set2 -> - let set1 = Value.to_ext Value.val_free set1 in - let set2 = Value.to_ext Value.val_free set2 in +let () = define2 "fresh_free_union" (repr_ext val_free) (repr_ext val_free) begin fun _ set1 set2 -> let ans = Id.Set.union set1 set2 in return (Value.of_ext Value.val_free ans) end -let () = define1 "fresh_free_of_ids" begin fun _ ids -> - let ids = Value.to_list Value.to_ident ids in +let () = define1 "fresh_free_of_ids" (list ident) begin fun _ ids -> let free = List.fold_right Id.Set.add ids Id.Set.empty in return (Value.of_ext Value.val_free free) end -let () = define1 "fresh_free_of_constr" begin fun _ c -> - let c = Value.to_constr c in +let () = define1 "fresh_free_of_constr" constr begin fun _ c -> Proofview.tclEVARMAP >>= fun sigma -> let rec fold accu c = match EConstr.kind sigma c with | Constr.Var id -> Id.Set.add id accu @@ -794,9 +737,7 @@ let () = define1 "fresh_free_of_constr" begin fun _ c -> return (Value.of_ext Value.val_free ans) end -let () = define2 "fresh_fresh" begin fun _ avoid id -> - let avoid = Value.to_ext Value.val_free avoid in - let id = Value.to_ident id in +let () = define2 "fresh_fresh" (repr_ext val_free) ident begin fun _ avoid id -> let nid = Namegen.next_ident_away_from id (fun id -> Id.Set.mem id avoid) in return (Value.of_ident nid) end @@ -999,7 +940,7 @@ let add_scope s f = let rec pr_scope = function | SexprStr (_, s) -> qstring s -| SexprInt (_, n) -> int n +| SexprInt (_, n) -> Pp.int n | SexprRec (_, (_, na), args) -> let na = match na with | None -> str "_" diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index c3f535c1bc..3e3dd362f7 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -49,6 +49,12 @@ exception LtacError of KerName.t * valexpr array * backtrace (** Conversion functions *) +let valexpr = { + r_of = (fun obj -> obj); + r_to = (fun obj -> obj); + r_id = true; +} + let of_unit () = ValInt 0 let to_unit = function @@ -216,6 +222,12 @@ let array r = { r_id = false; } +let block = { + r_of = of_tuple; + r_to = to_tuple; + r_id = false; +} + let of_constant c = of_ext val_constant c let to_constant c = to_ext val_constant c let constant = repr_ext val_constant diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index fe813b0e35..1789a8932f 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -64,6 +64,8 @@ val of_closure : closure -> valexpr val to_closure : valexpr -> closure val closure : closure repr +val block : valexpr array repr + val of_array : ('a -> valexpr) -> 'a array -> valexpr val to_array : (valexpr -> 'a) -> valexpr -> 'a array val array : 'a repr -> 'a array repr @@ -95,6 +97,8 @@ val of_ext : 'a Val.tag -> 'a -> valexpr val to_ext : 'a Val.tag -> valexpr -> 'a val repr_ext : 'a Val.tag -> 'a repr +val valexpr : valexpr repr + (** {5 Dynamic tags} *) val val_constr : EConstr.t Val.tag -- cgit v1.2.3 From 64a6ac3759b5d0ea635ff284606541b05c696996 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 6 Sep 2017 22:47:44 +0200 Subject: Using higher-order representation for closures. --- src/tac2core.ml | 31 ++++++++++----------------- src/tac2entries.ml | 6 +++--- src/tac2env.ml | 22 +------------------ src/tac2env.mli | 2 +- src/tac2expr.mli | 29 ++++++++----------------- src/tac2ffi.ml | 2 -- src/tac2ffi.mli | 7 +++--- src/tac2intern.ml | 2 +- src/tac2interp.ml | 63 ++++++++++++++++++++++++++++++++++++++++++++---------- src/tac2interp.mli | 2 +- src/tac2stdlib.ml | 2 +- 11 files changed, 83 insertions(+), 85 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 2f349e32af..603ddeecfd 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -100,7 +100,7 @@ let err_matchfailure bt = (** Helper functions *) -let thaw bt f = Tac2interp.interp_app bt f [v_unit] +let thaw bt f = f bt [v_unit] let throw bt e = Proofview.tclLIFT (Proofview.NonLogical.raise (e bt)) @@ -569,7 +569,7 @@ end (** (unit -> 'a) -> (exn -> 'a) -> 'a *) let () = define2 "plus" closure closure begin fun bt x k -> - Proofview.tclOR (thaw bt x) (fun e -> Tac2interp.interp_app bt k [Value.of_exn e]) + Proofview.tclOR (thaw bt x) (fun e -> k bt [Value.of_exn e]) end (** (unit -> 'a) -> 'a *) @@ -597,31 +597,22 @@ let () = define1 "enter" closure begin fun bt f -> Proofview.tclINDEPENDENT f >>= fun () -> return v_unit end -let k_var = Id.of_string "k" -let e_var = Id.of_string "e" -let prm_apply_kont_h = pname "apply_kont" - (** (unit -> 'a) -> ('a * ('exn -> 'a)) result *) let () = define1 "case" closure begin fun bt f -> Proofview.tclCASE (thaw bt f) >>= begin function | Proofview.Next (x, k) -> - let k = { - clos_ref = None; - clos_env = Id.Map.singleton k_var (Value.of_ext Value.val_kont k); - clos_var = [Name e_var]; - clos_exp = GTacPrm (prm_apply_kont_h, [GTacVar k_var; GTacVar e_var]); - } in - return (ValBlk (0, [| Value.of_tuple [| x; ValCls k |] |])) + let k bt = function + | [e] -> + let (e, info) = Value.to_exn e in + let e = set_bt bt e in + k (e, info) + | _ -> assert false + in + return (ValBlk (0, [| Value.of_tuple [| x; Value.of_closure k |] |])) | Proofview.Fail e -> return (ValBlk (1, [| Value.of_exn e |])) end end -(** 'a kont -> exn -> 'a *) -let () = define2 "apply_kont" (repr_ext val_kont) exn begin fun bt k (e, info) -> - let e = set_bt bt e in - k (e, info) -end - (** int -> int -> (unit -> 'a) -> 'a *) let () = define3 "focus" int int closure begin fun bt i j tac -> Proofview.tclFOCUS i j (thaw bt tac) @@ -695,7 +686,7 @@ let () = define2 "with_holes" closure closure begin fun bt x f -> thaw bt x >>= fun ans -> Proofview.tclEVARMAP >>= fun sigma -> Proofview.Unsafe.tclEVARS sigma0 >>= fun () -> - Tacticals.New.tclWITHHOLES false (Tac2interp.interp_app bt f [ans]) sigma + Tacticals.New.tclWITHHOLES false (f bt [ans]) sigma end let () = define1 "progress" closure begin fun bt f -> diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 9fd03ff5aa..97e1fe8e8e 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -686,7 +686,7 @@ type redefinition = { let perform_redefinition (_, redef) = let kn = redef.redef_kn in - let data, _ = Tac2env.interp_global kn in + let data = Tac2env.interp_global kn in let data = { data with Tac2env.gdata_expr = redef.redef_body } in Tac2env.define_global kn data @@ -715,7 +715,7 @@ let register_redefinition ?(local = false) (loc, qid) e = | TacAlias _ -> user_err ?loc (str "Cannot redefine syntactic abbreviations") in - let (data, _) = Tac2env.interp_global kn in + let data = Tac2env.interp_global kn in let () = if not (data.Tac2env.gdata_mutable) then user_err ?loc (str "The tactic " ++ pr_qualid qid ++ str " is not declared as mutable") @@ -800,7 +800,7 @@ let print_ltac ref = in match kn with | TacConstant kn -> - let data, _ = Tac2env.interp_global kn in + let data = Tac2env.interp_global kn in let e = data.Tac2env.gdata_expr in let (_, t) = data.Tac2env.gdata_type in let name = int_name () in diff --git a/src/tac2env.ml b/src/tac2env.ml index c04eaf7b0c..ef2b44afb9 100644 --- a/src/tac2env.ml +++ b/src/tac2env.ml @@ -52,33 +52,13 @@ let empty_state = { let ltac_state = Summary.ref empty_state ~name:"ltac2-state" -(** Get a dynamic value from syntactical value *) -let rec eval_pure kn = function -| GTacAtm (AtmInt n) -> ValInt n -| GTacRef kn -> - let { gdata_expr = e } = - try KNmap.find kn ltac_state.contents.ltac_tactics - with Not_found -> assert false - in - eval_pure (Some kn) e -| GTacFun (na, e) -> - ValCls { clos_ref = kn; clos_env = Id.Map.empty; clos_var = na; clos_exp = e } -| GTacCst (_, n, []) -> ValInt n -| GTacCst (_, n, el) -> ValBlk (n, Array.map_of_list eval_unnamed el) -| GTacOpn (kn, el) -> ValOpn (kn, Array.map_of_list eval_unnamed el) -| GTacAtm (AtmStr _) | GTacLet _ | GTacVar _ | GTacSet _ -| GTacApp _ | GTacCse _ | GTacPrj _ | GTacPrm _ | GTacExt _ | GTacWth _ -> - anomaly (Pp.str "Term is not a syntactical value") - -and eval_unnamed e = eval_pure None e - let define_global kn e = let state = !ltac_state in ltac_state := { state with ltac_tactics = KNmap.add kn e state.ltac_tactics } let interp_global kn = let data = KNmap.find kn ltac_state.contents.ltac_tactics in - (data, eval_pure (Some kn) data.gdata_expr) + data let define_constructor kn t = let state = !ltac_state in diff --git a/src/tac2env.mli b/src/tac2env.mli index e40958e1a0..49c9910a44 100644 --- a/src/tac2env.mli +++ b/src/tac2env.mli @@ -24,7 +24,7 @@ type global_data = { } val define_global : ltac_constant -> global_data -> unit -val interp_global : ltac_constant -> global_data * valexpr +val interp_global : ltac_constant -> global_data (** {5 Toplevel definition of types} *) diff --git a/src/tac2expr.mli b/src/tac2expr.mli index 36c3fbbe59..1b4a704b11 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -180,6 +180,13 @@ type strexpr = type tag = int +type frame = +| FrLtac of ltac_constant option +| FrPrim of ml_tactic_name +| FrExtn : ('a, 'b) Tac2dyn.Arg.tag * 'b -> frame + +type backtrace = frame list + type valexpr = | ValInt of int (** Immediate integers *) @@ -187,32 +194,14 @@ type valexpr = (** Structured blocks *) | ValStr of Bytes.t (** Strings *) -| ValCls of closure +| ValCls of ml_tactic (** Closures *) | ValOpn of KerName.t * valexpr array (** Open constructors *) | ValExt : 'a Tac2dyn.Val.tag * 'a -> valexpr (** Arbitrary data *) -and closure = { - mutable clos_env : valexpr Id.Map.t; - (** Mutable so that we can implement recursive functions imperatively *) - clos_var : Name.t list; - (** Bound variables *) - clos_exp : glb_tacexpr; - (** Body *) - clos_ref : ltac_constant option; - (** Global constant from which the closure originates *) -} - -type frame = -| FrLtac of ltac_constant option -| FrPrim of ml_tactic_name -| FrExtn : ('a, 'b) Tac2dyn.Arg.tag * 'b -> frame - -type backtrace = frame list - -type ml_tactic = backtrace -> valexpr list -> valexpr Proofview.tactic +and ml_tactic = backtrace -> valexpr list -> valexpr Proofview.tactic type environment = { env_ist : valexpr Id.Map.t; diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index 3e3dd362f7..fb97177c4d 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -34,8 +34,6 @@ let val_constructor = Val.create "constructor" let val_projection = Val.create "projection" let val_case = Val.create "case" let val_univ = Val.create "universe" -let val_kont : (Exninfo.iexn -> valexpr Proofview.tactic) Val.tag = - Val.create "kont" let val_free : Names.Id.Set.t Val.tag = Val.create "free" let extract_val (type a) (type b) (tag : a Val.tag) (tag' : b Val.tag) (v : b) : a = diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index 1789a8932f..dfc87f7db3 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -60,9 +60,9 @@ val of_ident : Id.t -> valexpr val to_ident : valexpr -> Id.t val ident : Id.t repr -val of_closure : closure -> valexpr -val to_closure : valexpr -> closure -val closure : closure repr +val of_closure : ml_tactic -> valexpr +val to_closure : valexpr -> ml_tactic +val closure : ml_tactic repr val block : valexpr array repr @@ -113,7 +113,6 @@ val val_constructor : constructor Val.tag val val_projection : Projection.t Val.tag val val_case : Constr.case_info Val.tag val val_univ : Univ.universe_level Val.tag -val val_kont : (Exninfo.iexn -> valexpr Proofview.tactic) Val.tag val val_free : Id.Set.t Val.tag val val_exn : Exninfo.iexn Tac2dyn.Val.tag diff --git a/src/tac2intern.ml b/src/tac2intern.ml index d1a3e13cdb..2dcd8b8da3 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -654,7 +654,7 @@ let rec intern_rec env (loc, e) = match e with let sch = Id.Map.find id env.env_var in (GTacVar id, fresh_mix_type_scheme env sch) | ArgArg (TacConstant kn) -> - let { Tac2env.gdata_type = sch }, _ = + let { Tac2env.gdata_type = sch } = try Tac2env.interp_global kn with Not_found -> CErrors.anomaly (str "Missing hardwired primitive " ++ KerName.print kn) diff --git a/src/tac2interp.ml b/src/tac2interp.ml index 6bee5a0794..08cebc0af0 100644 --- a/src/tac2interp.ml +++ b/src/tac2interp.ml @@ -21,6 +21,17 @@ let empty_environment = { env_bkt = []; } +type closure = { + mutable clos_env : valexpr Id.Map.t; + (** Mutable so that we can implement recursive functions imperatively *) + clos_var : Name.t list; + (** Bound variables *) + clos_exp : glb_tacexpr; + (** Body *) + clos_ref : ltac_constant option; + (** Global constant from which the closure originates *) +} + let push_name ist id v = match id with | Anonymous -> ist | Name id -> { ist with env_ist = Id.Map.add id v ist.env_ist } @@ -30,7 +41,10 @@ let get_var ist id = anomaly (str "Unbound variable " ++ Id.print id) let get_ref ist kn = - try snd (Tac2env.interp_global kn) with Not_found -> + try + let data = Tac2env.interp_global kn in + data.Tac2env.gdata_expr + with Not_found -> anomaly (str "Unbound reference" ++ KerName.print kn) let return = Proofview.tclUNIT @@ -39,14 +53,17 @@ let rec interp (ist : environment) = function | GTacAtm (AtmInt n) -> return (ValInt n) | GTacAtm (AtmStr s) -> return (ValStr (Bytes.of_string s)) | GTacVar id -> return (get_var ist id) -| GTacRef qid -> return (get_ref ist qid) +| GTacRef kn -> + let data = get_ref ist kn in + return (eval_pure (Some kn) data) | GTacFun (ids, e) -> let cls = { clos_ref = None; clos_env = ist.env_ist; clos_var = ids; clos_exp = e } in - return (ValCls cls) + let f = interp_app cls in + return (ValCls f) | GTacApp (f, args) -> interp ist f >>= fun f -> Proofview.Monad.List.map (fun e -> interp ist e) args >>= fun args -> - interp_app ist.env_bkt (Tac2ffi.to_closure f) args + Tac2ffi.to_closure f ist.env_bkt args | GTacLet (false, el, e) -> let fold accu (na, e) = interp ist e >>= fun e -> @@ -58,17 +75,18 @@ let rec interp (ist : environment) = function let map (na, e) = match e with | GTacFun (ids, e) -> let cls = { clos_ref = None; clos_env = ist.env_ist; clos_var = ids; clos_exp = e } in - na, cls + let f = ValCls (interp_app cls) in + na, cls, f | _ -> anomaly (str "Ill-formed recursive function") in let fixs = List.map map el in - let fold accu (na, cls) = match na with + let fold accu (na, _, cls) = match na with | Anonymous -> accu - | Name id -> { ist with env_ist = Id.Map.add id (ValCls cls) accu.env_ist } + | Name id -> { ist with env_ist = Id.Map.add id cls accu.env_ist } in let ist = List.fold_left fold ist fixs in (** Hack to make a cycle imperatively in the environment *) - let iter (_, e) = e.clos_env <- ist.env_ist in + let iter (_, e, _) = e.clos_env <- ist.env_ist in let () = List.iter iter fixs in interp ist e | GTacCst (_, n, []) -> return (ValInt n) @@ -96,15 +114,16 @@ let rec interp (ist : environment) = function let ist = { ist with env_bkt = FrExtn (tag, e) :: ist.env_bkt } in tpe.Tac2env.ml_interp ist e -and interp_app bt f args = match f with +and interp_app f = (); fun bt args -> match f with | { clos_env = ist; clos_var = ids; clos_exp = e; clos_ref = kn } -> let rec push ist ids args = match ids, args with | [], [] -> interp ist e | [], _ :: _ -> - interp ist e >>= fun f -> interp_app bt (Tac2ffi.to_closure f) args + interp ist e >>= fun f -> Tac2ffi.to_closure f bt args | _ :: _, [] -> let cls = { clos_ref = kn; clos_env = ist.env_ist; clos_var = ids; clos_exp = e } in - return (ValCls cls) + let f = interp_app cls in + return (ValCls f) | id :: ids, arg :: args -> push (push_name ist id arg) ids args in push { env_ist = ist; env_bkt = FrLtac kn :: bt } ids args @@ -147,6 +166,28 @@ and interp_set ist e p r = match e with | ValInt _ | ValExt _ | ValStr _ | ValCls _ | ValOpn _ -> anomaly (str "Unexpected value shape") +and eval_pure kn = function +| GTacAtm (AtmInt n) -> ValInt n +| GTacRef kn -> + let { Tac2env.gdata_expr = e } = + try Tac2env.interp_global kn + with Not_found -> assert false + in + eval_pure (Some kn) e +| GTacFun (na, e) -> + let cls = { clos_ref = kn; clos_env = Id.Map.empty; clos_var = na; clos_exp = e } in + let f = interp_app cls in + ValCls f +| GTacCst (_, n, []) -> ValInt n +| GTacCst (_, n, el) -> ValBlk (n, Array.map_of_list eval_unnamed el) +| GTacOpn (kn, el) -> ValOpn (kn, Array.map_of_list eval_unnamed el) +| GTacAtm (AtmStr _) | GTacLet _ | GTacVar _ | GTacSet _ +| GTacApp _ | GTacCse _ | GTacPrj _ | GTacPrm _ | GTacExt _ | GTacWth _ -> + anomaly (Pp.str "Term is not a syntactical value") + +and eval_unnamed e = eval_pure None e + + (** Cross-boundary hacks. *) open Geninterp diff --git a/src/tac2interp.mli b/src/tac2interp.mli index ca263b2c4b..48af59b178 100644 --- a/src/tac2interp.mli +++ b/src/tac2interp.mli @@ -13,7 +13,7 @@ val empty_environment : environment val interp : environment -> glb_tacexpr -> valexpr Proofview.tactic -val interp_app : backtrace -> closure -> valexpr list -> valexpr Proofview.tactic +(* val interp_app : closure -> ml_tactic *) (** {5 Cross-boundary encodings} *) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 79af41b7d0..14ad8695ca 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -18,7 +18,7 @@ module Value = Tac2ffi let return x = Proofview.tclUNIT x let v_unit = Value.of_unit () -let thaw bt f = Tac2interp.interp_app bt (Value.to_closure f) [v_unit] +let thaw bt f = (Value.to_closure f) bt [v_unit] let to_pair f g = function | ValBlk (0, [| x; y |]) -> (f x, g y) -- cgit v1.2.3 From d6997e31e7fc4cfc6e020bf1ab53e6b1fa3f74fe Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 6 Sep 2017 23:46:41 +0200 Subject: Fix coq/ltac2#23: Int.compare should not be uniformly 0. Stupid typo. --- src/tac2core.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 603ddeecfd..7f136b48ae 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -230,7 +230,7 @@ let () = define2 "int_equal" int int begin fun _ m n -> end let binop n f = define2 n int int begin fun _ m n -> - return (Value.of_int (f m m)) + return (Value.of_int (f m n)) end let () = binop "int_compare" Int.compare -- cgit v1.2.3 From 2bea4137bd0841de7273a5adf9a72bd2e786fb68 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 6 Sep 2017 23:50:33 +0200 Subject: Communicate the backtrace through the monad. --- src/tac2core.ml | 233 +++++++++++++++++++++++++++------------------------- src/tac2core.mli | 2 +- src/tac2entries.ml | 8 +- src/tac2expr.mli | 3 +- src/tac2interp.ml | 46 ++++++++--- src/tac2interp.mli | 8 ++ src/tac2stdlib.ml | 221 +++++++++++++++++++++++++------------------------ src/tac2tactics.ml | 48 +++++------ src/tac2tactics.mli | 22 ++--- 9 files changed, 313 insertions(+), 278 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 7f136b48ae..4a35442b04 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -100,9 +100,14 @@ let err_matchfailure bt = (** Helper functions *) -let thaw bt f = f bt [v_unit] +let thaw f = f [v_unit] -let throw bt e = Proofview.tclLIFT (Proofview.NonLogical.raise (e bt)) +let throw e = + Tac2interp.get_backtrace >>= fun bt -> + Proofview.tclLIFT (Proofview.NonLogical.raise (e bt)) + +let fail e = + Tac2interp.get_backtrace >>= fun bt -> Proofview.tclZERO (e bt) let set_bt bt e = match e with | Tac2interp.LtacError (kn, args, _) -> Tac2interp.LtacError (kn, args, bt) @@ -117,13 +122,13 @@ let wrap f = let wrap_unit f = return () >>= fun () -> f (); return v_unit -let assert_focussed bt = +let assert_focussed = Proofview.Goal.goals >>= fun gls -> match gls with | [_] -> Proofview.tclUNIT () - | [] | _ :: _ :: _ -> throw bt err_notfocussed + | [] | _ :: _ :: _ -> throw err_notfocussed -let pf_apply bt f = +let pf_apply f = Proofview.Goal.goals >>= function | [] -> Proofview.tclENV >>= fun env -> @@ -133,103 +138,103 @@ let pf_apply bt f = gl >>= fun gl -> f (Proofview.Goal.env gl) (Tacmach.New.project gl) | _ :: _ :: _ -> - throw bt err_notfocussed + throw err_notfocussed (** Primitives *) -let define0 name f = Tac2env.define_primitive (pname name) begin fun bt arg -> match arg with -| [_] -> f bt +let define0 name f = Tac2env.define_primitive (pname name) begin fun arg -> match arg with +| [_] -> f | _ -> assert false end -let define1 name r0 f = Tac2env.define_primitive (pname name) begin fun bt arg -> match arg with -| [x] -> f bt (r0.Value.r_to x) +let define1 name r0 f = Tac2env.define_primitive (pname name) begin fun arg -> match arg with +| [x] -> f (r0.Value.r_to x) | _ -> assert false end -let define2 name r0 r1 f = Tac2env.define_primitive (pname name) begin fun bt arg -> match arg with -| [x; y] -> f bt (r0.Value.r_to x) (r1.Value.r_to y) +let define2 name r0 r1 f = Tac2env.define_primitive (pname name) begin fun arg -> match arg with +| [x; y] -> f (r0.Value.r_to x) (r1.Value.r_to y) | _ -> assert false end -let define3 name r0 r1 r2 f = Tac2env.define_primitive (pname name) begin fun bt arg -> match arg with -| [x; y; z] -> f bt (r0.Value.r_to x) (r1.Value.r_to y) (r2.Value.r_to z) +let define3 name r0 r1 r2 f = Tac2env.define_primitive (pname name) begin fun arg -> match arg with +| [x; y; z] -> f (r0.Value.r_to x) (r1.Value.r_to y) (r2.Value.r_to z) | _ -> assert false end (** Printing *) -let () = define1 "print" pp begin fun _ pp -> +let () = define1 "print" pp begin fun pp -> wrap_unit (fun () -> Feedback.msg_notice pp) end -let () = define1 "message_of_int" int begin fun _ n -> +let () = define1 "message_of_int" int begin fun n -> return (Value.of_pp (Pp.int n)) end -let () = define1 "message_of_string" string begin fun _ s -> +let () = define1 "message_of_string" string begin fun s -> return (Value.of_pp (str (Bytes.to_string s))) end -let () = define1 "message_of_constr" constr begin fun bt c -> - pf_apply bt begin fun env sigma -> +let () = define1 "message_of_constr" constr begin fun c -> + pf_apply begin fun env sigma -> let pp = Printer.pr_econstr_env env sigma c in return (Value.of_pp pp) end end -let () = define1 "message_of_ident" ident begin fun _ c -> +let () = define1 "message_of_ident" ident begin fun c -> let pp = Id.print c in return (Value.of_pp pp) end -let () = define2 "message_concat" pp pp begin fun _ m1 m2 -> +let () = define2 "message_concat" pp pp begin fun m1 m2 -> return (Value.of_pp (Pp.app m1 m2)) end (** Array *) -let () = define2 "array_make" int valexpr begin fun bt n x -> - if n < 0 || n > Sys.max_array_length then throw bt err_outofbounds +let () = define2 "array_make" int valexpr begin fun n x -> + if n < 0 || n > Sys.max_array_length then throw err_outofbounds else wrap (fun () -> ValBlk (0, Array.make n x)) end -let () = define1 "array_length" block begin fun _ v -> +let () = define1 "array_length" block begin fun v -> return (ValInt (Array.length v)) end -let () = define3 "array_set" block int valexpr begin fun bt v n x -> - if n < 0 || n >= Array.length v then throw bt err_outofbounds +let () = define3 "array_set" block int valexpr begin fun v n x -> + if n < 0 || n >= Array.length v then throw err_outofbounds else wrap_unit (fun () -> v.(n) <- x) end -let () = define2 "array_get" block int begin fun bt v n -> - if n < 0 || n >= Array.length v then throw bt err_outofbounds +let () = define2 "array_get" block int begin fun v n -> + if n < 0 || n >= Array.length v then throw err_outofbounds else wrap (fun () -> v.(n)) end (** Ident *) -let () = define2 "ident_equal" ident ident begin fun _ id1 id2 -> +let () = define2 "ident_equal" ident ident begin fun id1 id2 -> return (Value.of_bool (Id.equal id1 id2)) end -let () = define1 "ident_to_string" ident begin fun _ id -> +let () = define1 "ident_to_string" ident begin fun id -> return (Value.of_string (Id.to_string id)) end -let () = define1 "ident_of_string" string begin fun _ s -> +let () = define1 "ident_of_string" string begin fun s -> let id = try Some (Id.of_string s) with _ -> None in return (Value.of_option Value.of_ident id) end (** Int *) -let () = define2 "int_equal" int int begin fun _ m n -> +let () = define2 "int_equal" int int begin fun m n -> return (Value.of_bool (m == n)) end -let binop n f = define2 n int int begin fun _ m n -> +let binop n f = define2 n int int begin fun m n -> return (Value.of_int (f m n)) end @@ -238,52 +243,52 @@ let () = binop "int_add" (+) let () = binop "int_sub" (-) let () = binop "int_mul" ( * ) -let () = define1 "int_neg" int begin fun _ m -> +let () = define1 "int_neg" int begin fun m -> return (Value.of_int (~- m)) end (** String *) -let () = define2 "string_make" int char begin fun bt n c -> - if n < 0 || n > Sys.max_string_length then throw bt err_outofbounds +let () = define2 "string_make" int char begin fun n c -> + if n < 0 || n > Sys.max_string_length then throw err_outofbounds else wrap (fun () -> Value.of_string (Bytes.make n c)) end -let () = define1 "string_length" string begin fun _ s -> +let () = define1 "string_length" string begin fun s -> return (Value.of_int (Bytes.length s)) end -let () = define3 "string_set" string int char begin fun bt s n c -> - if n < 0 || n >= Bytes.length s then throw bt err_outofbounds +let () = define3 "string_set" string int char begin fun s n c -> + if n < 0 || n >= Bytes.length s then throw err_outofbounds else wrap_unit (fun () -> Bytes.set s n c) end -let () = define2 "string_get" string int begin fun bt s n -> - if n < 0 || n >= Bytes.length s then throw bt err_outofbounds +let () = define2 "string_get" string int begin fun s n -> + if n < 0 || n >= Bytes.length s then throw err_outofbounds else wrap (fun () -> Value.of_char (Bytes.get s n)) end (** Terms *) (** constr -> constr *) -let () = define1 "constr_type" constr begin fun bt c -> +let () = define1 "constr_type" constr begin fun c -> let get_type env sigma = Proofview.V82.wrap_exceptions begin fun () -> let (sigma, t) = Typing.type_of env sigma c in let t = Value.of_constr t in Proofview.Unsafe.tclEVARS sigma <*> Proofview.tclUNIT t end in - pf_apply bt get_type + pf_apply get_type end (** constr -> constr *) -let () = define2 "constr_equal" constr constr begin fun _ c1 c2 -> +let () = define2 "constr_equal" constr constr begin fun c1 c2 -> Proofview.tclEVARMAP >>= fun sigma -> let b = EConstr.eq_constr sigma c1 c2 in Proofview.tclUNIT (Value.of_bool b) end -let () = define1 "constr_kind" constr begin fun _ c -> +let () = define1 "constr_kind" constr begin fun c -> let open Constr in Proofview.tclEVARMAP >>= fun sigma -> return begin match EConstr.kind sigma c with @@ -377,7 +382,7 @@ let () = define1 "constr_kind" constr begin fun _ c -> end end -let () = define1 "constr_make" valexpr begin fun _ knd -> +let () = define1 "constr_make" valexpr begin fun knd -> let open Constr in let c = match knd with | ValBlk (0, [|n|]) -> @@ -457,8 +462,8 @@ let () = define1 "constr_make" valexpr begin fun _ knd -> return (Value.of_constr c) end -let () = define1 "constr_check" constr begin fun bt c -> - pf_apply bt begin fun env sigma -> +let () = define1 "constr_check" constr begin fun c -> + pf_apply begin fun env sigma -> try let (sigma, _) = Typing.type_of env sigma c in Proofview.Unsafe.tclEVARS sigma >>= fun () -> @@ -469,26 +474,26 @@ let () = define1 "constr_check" constr begin fun bt c -> end end -let () = define3 "constr_substnl" (list constr) int constr begin fun _ subst k c -> +let () = define3 "constr_substnl" (list constr) int constr begin fun subst k c -> let ans = EConstr.Vars.substnl subst k c in return (Value.of_constr ans) end -let () = define3 "constr_closenl" (list ident) int constr begin fun _ ids k c -> +let () = define3 "constr_closenl" (list ident) int constr begin fun ids k c -> let ans = EConstr.Vars.substn_vars k ids c in return (Value.of_constr ans) end (** Patterns *) -let () = define2 "pattern_matches" pattern constr begin fun bt pat c -> - pf_apply bt begin fun env sigma -> +let () = define2 "pattern_matches" pattern constr begin fun pat c -> + pf_apply begin fun env sigma -> let ans = try Some (Constr_matching.matches env sigma pat c) with Constr_matching.PatternMatchingFailure -> None in begin match ans with - | None -> Proofview.tclZERO (err_matchfailure bt) + | None -> fail err_matchfailure | Some ans -> let ans = Id.Map.bindings ans in let of_pair (id, c) = Value.of_tuple [| Value.of_ident id; Value.of_constr c |] in @@ -497,30 +502,30 @@ let () = define2 "pattern_matches" pattern constr begin fun bt pat c -> end end -let () = define2 "pattern_matches_subterm" pattern constr begin fun bt pat c -> +let () = define2 "pattern_matches_subterm" pattern constr begin fun pat c -> let open Constr_matching in let rec of_ans s = match IStream.peek s with - | IStream.Nil -> Proofview.tclZERO (err_matchfailure bt) + | IStream.Nil -> fail err_matchfailure | IStream.Cons ({ m_sub = (_, sub); m_ctx }, s) -> let ans = Id.Map.bindings sub in let of_pair (id, c) = Value.of_tuple [| Value.of_ident id; Value.of_constr c |] in let ans = Value.of_tuple [| Value.of_constr m_ctx; Value.of_list of_pair ans |] in Proofview.tclOR (return ans) (fun _ -> of_ans s) in - pf_apply bt begin fun env sigma -> + pf_apply begin fun env sigma -> let ans = Constr_matching.match_appsubterm env sigma pat c in of_ans ans end end -let () = define2 "pattern_matches_vect" pattern constr begin fun bt pat c -> - pf_apply bt begin fun env sigma -> +let () = define2 "pattern_matches_vect" pattern constr begin fun pat c -> + pf_apply begin fun env sigma -> let ans = try Some (Constr_matching.matches env sigma pat c) with Constr_matching.PatternMatchingFailure -> None in begin match ans with - | None -> Proofview.tclZERO (err_matchfailure bt) + | None -> fail err_matchfailure | Some ans -> let ans = Id.Map.bindings ans in let ans = Array.map_of_list snd ans in @@ -529,23 +534,23 @@ let () = define2 "pattern_matches_vect" pattern constr begin fun bt pat c -> end end -let () = define2 "pattern_matches_subterm_vect" pattern constr begin fun bt pat c -> +let () = define2 "pattern_matches_subterm_vect" pattern constr begin fun pat c -> let open Constr_matching in let rec of_ans s = match IStream.peek s with - | IStream.Nil -> Proofview.tclZERO (err_matchfailure bt) + | IStream.Nil -> fail err_matchfailure | IStream.Cons ({ m_sub = (_, sub); m_ctx }, s) -> let ans = Id.Map.bindings sub in let ans = Array.map_of_list snd ans in let ans = Value.of_tuple [| Value.of_constr m_ctx; Value.of_array Value.of_constr ans |] in Proofview.tclOR (return ans) (fun _ -> of_ans s) in - pf_apply bt begin fun env sigma -> + pf_apply begin fun env sigma -> let ans = Constr_matching.match_appsubterm env sigma pat c in of_ans ans end end -let () = define2 "pattern_instantiate" constr constr begin fun _ ctx c -> +let () = define2 "pattern_instantiate" constr constr begin fun ctx c -> let ctx = EConstr.Unsafe.to_constr ctx in let c = EConstr.Unsafe.to_constr c in let ans = Termops.subst_meta [Constr_matching.special_meta, c] ctx in @@ -554,7 +559,8 @@ end (** Error *) -let () = define1 "throw" exn begin fun bt (e, info) -> +let () = define1 "throw" exn begin fun (e, info) -> + Tac2interp.get_backtrace >>= fun bt -> let e = set_bt bt e in Proofview.tclLIFT (Proofview.NonLogical.raise ~info e) end @@ -562,48 +568,50 @@ end (** Control *) (** exn -> 'a *) -let () = define1 "zero" exn begin fun bt (e, info) -> +let () = define1 "zero" exn begin fun (e, info) -> + Tac2interp.get_backtrace >>= fun bt -> let e = set_bt bt e in Proofview.tclZERO ~info e end (** (unit -> 'a) -> (exn -> 'a) -> 'a *) -let () = define2 "plus" closure closure begin fun bt x k -> - Proofview.tclOR (thaw bt x) (fun e -> k bt [Value.of_exn e]) +let () = define2 "plus" closure closure begin fun x k -> + Proofview.tclOR (thaw x) (fun e -> k [Value.of_exn e]) end (** (unit -> 'a) -> 'a *) -let () = define1 "once" closure begin fun bt f -> - Proofview.tclONCE (thaw bt f) +let () = define1 "once" closure begin fun f -> + Proofview.tclONCE (thaw f) end (** (unit -> unit) list -> unit *) -let () = define1 "dispatch" (list closure) begin fun bt l -> - let l = List.map (fun f -> Proofview.tclIGNORE (thaw bt f)) l in +let () = define1 "dispatch" (list closure) begin fun l -> + let l = List.map (fun f -> Proofview.tclIGNORE (thaw f)) l in Proofview.tclDISPATCH l >>= fun () -> return v_unit end (** (unit -> unit) list -> (unit -> unit) -> (unit -> unit) list -> unit *) -let () = define3 "extend" (list closure) closure (list closure) begin fun bt lft tac rgt -> - let lft = List.map (fun f -> Proofview.tclIGNORE (thaw bt f)) lft in - let tac = Proofview.tclIGNORE (thaw bt tac) in - let rgt = List.map (fun f -> Proofview.tclIGNORE (thaw bt f)) rgt in +let () = define3 "extend" (list closure) closure (list closure) begin fun lft tac rgt -> + let lft = List.map (fun f -> Proofview.tclIGNORE (thaw f)) lft in + let tac = Proofview.tclIGNORE (thaw tac) in + let rgt = List.map (fun f -> Proofview.tclIGNORE (thaw f)) rgt in Proofview.tclEXTEND lft tac rgt >>= fun () -> return v_unit end (** (unit -> unit) -> unit *) -let () = define1 "enter" closure begin fun bt f -> - let f = Proofview.tclIGNORE (thaw bt f) in +let () = define1 "enter" closure begin fun f -> + let f = Proofview.tclIGNORE (thaw f) in Proofview.tclINDEPENDENT f >>= fun () -> return v_unit end (** (unit -> 'a) -> ('a * ('exn -> 'a)) result *) -let () = define1 "case" closure begin fun bt f -> - Proofview.tclCASE (thaw bt f) >>= begin function +let () = define1 "case" closure begin fun f -> + Proofview.tclCASE (thaw f) >>= begin function | Proofview.Next (x, k) -> - let k bt = function + let k = function | [e] -> let (e, info) = Value.to_exn e in + Tac2interp.get_backtrace >>= fun bt -> let e = set_bt bt e in k (e, info) | _ -> assert false @@ -614,31 +622,31 @@ let () = define1 "case" closure begin fun bt f -> end (** int -> int -> (unit -> 'a) -> 'a *) -let () = define3 "focus" int int closure begin fun bt i j tac -> - Proofview.tclFOCUS i j (thaw bt tac) +let () = define3 "focus" int int closure begin fun i j tac -> + Proofview.tclFOCUS i j (thaw tac) end (** unit -> unit *) -let () = define0 "shelve" begin fun _ -> +let () = define0 "shelve" begin Proofview.shelve >>= fun () -> return v_unit end (** unit -> unit *) -let () = define0 "shelve_unifiable" begin fun _ -> +let () = define0 "shelve_unifiable" begin Proofview.shelve_unifiable >>= fun () -> return v_unit end -let () = define1 "new_goal" int begin fun bt ev -> +let () = define1 "new_goal" int begin fun ev -> let ev = Evar.unsafe_of_int ev in Proofview.tclEVARMAP >>= fun sigma -> if Evd.mem sigma ev then Proofview.Unsafe.tclNEWGOALS [ev] <*> Proofview.tclUNIT v_unit - else throw bt err_notfound + else throw err_notfound end (** unit -> constr *) -let () = define0 "goal" begin fun bt -> - assert_focussed bt >>= fun () -> +let () = define0 "goal" begin + assert_focussed >>= fun () -> Proofview.Goal.enter_one begin fun gl -> let concl = Tacmach.New.pf_nf_concl gl in return (Value.of_constr concl) @@ -646,8 +654,8 @@ let () = define0 "goal" begin fun bt -> end (** ident -> constr *) -let () = define1 "hyp" ident begin fun bt id -> - pf_apply bt begin fun env _ -> +let () = define1 "hyp" ident begin fun id -> + pf_apply begin fun env _ -> let mem = try ignore (Environ.lookup_named id env); true with Not_found -> false in if mem then return (Value.of_constr (EConstr.mkVar id)) else Tacticals.New.tclZEROMSG @@ -655,8 +663,8 @@ let () = define1 "hyp" ident begin fun bt id -> end end -let () = define0 "hyps" begin fun bt -> - pf_apply bt begin fun env _ -> +let () = define0 "hyps" begin + pf_apply begin fun env _ -> let open Context.Named.Declaration in let hyps = List.rev (Environ.named_context env) in let map = function @@ -673,52 +681,52 @@ let () = define0 "hyps" begin fun bt -> end (** (unit -> constr) -> unit *) -let () = define1 "refine" closure begin fun bt c -> - let c = thaw bt c >>= fun c -> Proofview.tclUNIT ((), Value.to_constr c) in +let () = define1 "refine" closure begin fun c -> + let c = thaw c >>= fun c -> Proofview.tclUNIT ((), Value.to_constr c) in Proofview.Goal.enter begin fun gl -> let gl = Proofview.Goal.assume gl in Refine.generic_refine ~typecheck:true c gl end >>= fun () -> return v_unit end -let () = define2 "with_holes" closure closure begin fun bt x f -> +let () = define2 "with_holes" closure closure begin fun x f -> Proofview.tclEVARMAP >>= fun sigma0 -> - thaw bt x >>= fun ans -> + thaw x >>= fun ans -> Proofview.tclEVARMAP >>= fun sigma -> Proofview.Unsafe.tclEVARS sigma0 >>= fun () -> - Tacticals.New.tclWITHHOLES false (f bt [ans]) sigma + Tacticals.New.tclWITHHOLES false (f [ans]) sigma end -let () = define1 "progress" closure begin fun bt f -> - Proofview.tclPROGRESS (thaw bt f) +let () = define1 "progress" closure begin fun f -> + Proofview.tclPROGRESS (thaw f) end -let () = define2 "abstract" (option ident) closure begin fun bt id f -> - Tactics.tclABSTRACT id (Proofview.tclIGNORE (thaw bt f)) >>= fun () -> +let () = define2 "abstract" (option ident) closure begin fun id f -> + Tactics.tclABSTRACT id (Proofview.tclIGNORE (thaw f)) >>= fun () -> return v_unit end -let () = define2 "time" (option string) closure begin fun bt s f -> - Proofview.tclTIME s (thaw bt f) +let () = define2 "time" (option string) closure begin fun s f -> + Proofview.tclTIME s (thaw f) end -let () = define0 "check_interrupt" begin fun bt -> +let () = define0 "check_interrupt" begin Proofview.tclCHECKINTERRUPT >>= fun () -> return v_unit end (** Fresh *) -let () = define2 "fresh_free_union" (repr_ext val_free) (repr_ext val_free) begin fun _ set1 set2 -> +let () = define2 "fresh_free_union" (repr_ext val_free) (repr_ext val_free) begin fun set1 set2 -> let ans = Id.Set.union set1 set2 in return (Value.of_ext Value.val_free ans) end -let () = define1 "fresh_free_of_ids" (list ident) begin fun _ ids -> +let () = define1 "fresh_free_of_ids" (list ident) begin fun ids -> let free = List.fold_right Id.Set.add ids Id.Set.empty in return (Value.of_ext Value.val_free free) end -let () = define1 "fresh_free_of_constr" constr begin fun _ c -> +let () = define1 "fresh_free_of_constr" constr begin fun c -> Proofview.tclEVARMAP >>= fun sigma -> let rec fold accu c = match EConstr.kind sigma c with | Constr.Var id -> Id.Set.add id accu @@ -728,7 +736,7 @@ let () = define1 "fresh_free_of_constr" constr begin fun _ c -> return (Value.of_ext Value.val_free ans) end -let () = define2 "fresh_fresh" (repr_ext val_free) ident begin fun _ avoid id -> +let () = define2 "fresh_fresh" (repr_ext val_free) ident begin fun avoid id -> let nid = Namegen.next_ident_away_from id (fun id -> Id.Set.mem id avoid) in return (Value.of_ident nid) end @@ -769,9 +777,8 @@ let intern_constr self ist c = let interp_constr flags ist c = let open Pretyping in - let bt = ist.env_bkt in let ist = to_lvar ist in - pf_apply bt begin fun env sigma -> + pf_apply begin fun env sigma -> Proofview.V82.wrap_exceptions begin fun () -> let (sigma, c) = understand_ltac flags env sigma ist WithoutTypeConstraint c in let c = ValExt (Value.val_constr, c) in @@ -865,7 +872,7 @@ let () = GlbVal tac, gtypref t_unit in let interp ist tac = - let ist = { ist with env_ist = Id.Map.empty } in + let ist = { env_ist = Id.Map.empty } in let lfun = Tac2interp.set_env ist Id.Map.empty in let ist = Ltac_plugin.Tacinterp.default_ist () in (** FUCK YOU API *) @@ -911,8 +918,8 @@ let () = (** FUCK YOU API *) let idtac = (Obj.magic idtac : Geninterp.Val.t) in let interp ist tac = - let ist = Tac2interp.get_env ist.Geninterp.lfun in - let ist = { ist with env_ist = Id.Map.empty } in +(* let ist = Tac2interp.get_env ist.Geninterp.lfun in *) + let ist = { env_ist = Id.Map.empty } in Tac2interp.interp ist tac >>= fun _ -> Ftactic.return idtac in diff --git a/src/tac2core.mli b/src/tac2core.mli index b5800a7172..9fae65bb3e 100644 --- a/src/tac2core.mli +++ b/src/tac2core.mli @@ -27,4 +27,4 @@ val c_false : ltac_constructor end -val pf_apply : backtrace -> (Environ.env -> Evd.evar_map -> 'a Proofview.tactic) -> 'a Proofview.tactic +val pf_apply : (Environ.env -> Evd.evar_map -> 'a Proofview.tactic) -> 'a Proofview.tactic diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 97e1fe8e8e..eed7eb6def 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -748,14 +748,12 @@ let register_struct ?local str = match str with (** Toplevel exception *) -let print_ltac2_backtrace = ref false - let _ = Goptions.declare_bool_option { Goptions.optdepr = false; Goptions.optname = "print Ltac2 backtrace"; Goptions.optkey = ["Ltac2"; "Backtrace"]; - Goptions.optread = (fun () -> !print_ltac2_backtrace); - Goptions.optwrite = (fun b -> print_ltac2_backtrace := b); + Goptions.optread = (fun () -> !Tac2interp.print_ltac2_backtrace); + Goptions.optwrite = (fun b -> Tac2interp.print_ltac2_backtrace := b); } let pr_frame = function @@ -773,7 +771,7 @@ let () = register_handler begin function | Tac2interp.LtacError (kn, _, bt) -> let c = Tac2print.pr_constructor kn in (** FIXME *) let bt = - if !print_ltac2_backtrace then + if !Tac2interp.print_ltac2_backtrace then fnl () ++ str "Backtrace:" ++ fnl () ++ prlist_with_sep fnl pr_frame bt else mt () diff --git a/src/tac2expr.mli b/src/tac2expr.mli index 1b4a704b11..77e2cfef0e 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -201,9 +201,8 @@ type valexpr = | ValExt : 'a Tac2dyn.Val.tag * 'a -> valexpr (** Arbitrary data *) -and ml_tactic = backtrace -> valexpr list -> valexpr Proofview.tactic +and ml_tactic = valexpr list -> valexpr Proofview.tactic type environment = { env_ist : valexpr Id.Map.t; - env_bkt : backtrace; } diff --git a/src/tac2interp.ml b/src/tac2interp.ml index 08cebc0af0..f37b4f8e9c 100644 --- a/src/tac2interp.ml +++ b/src/tac2interp.ml @@ -16,9 +16,34 @@ open Tac2expr exception LtacError = Tac2ffi.LtacError +let backtrace : backtrace Evd.Store.field = Evd.Store.field () + +let print_ltac2_backtrace = ref false + +let get_backtrace = + Proofview.tclEVARMAP >>= fun sigma -> + match Evd.Store.get (Evd.get_extra_data sigma) backtrace with + | None -> Proofview.tclUNIT [] + | Some bt -> Proofview.tclUNIT bt + +let set_backtrace bt = + Proofview.tclEVARMAP >>= fun sigma -> + let store = Evd.get_extra_data sigma in + let store = Evd.Store.set store backtrace bt in + let sigma = Evd.set_extra_data store sigma in + Proofview.Unsafe.tclEVARS sigma + +let with_frame frame tac = + if !print_ltac2_backtrace then + get_backtrace >>= fun bt -> + set_backtrace (frame :: bt) >>= fun () -> + tac >>= fun ans -> + set_backtrace bt >>= fun () -> + Proofview.tclUNIT ans + else tac + let empty_environment = { env_ist = Id.Map.empty; - env_bkt = []; } type closure = { @@ -34,7 +59,7 @@ type closure = { let push_name ist id v = match id with | Anonymous -> ist -| Name id -> { ist with env_ist = Id.Map.add id v ist.env_ist } +| Name id -> { env_ist = Id.Map.add id v ist.env_ist } let get_var ist id = try Id.Map.find id ist.env_ist with Not_found -> @@ -63,7 +88,7 @@ let rec interp (ist : environment) = function | GTacApp (f, args) -> interp ist f >>= fun f -> Proofview.Monad.List.map (fun e -> interp ist e) args >>= fun args -> - Tac2ffi.to_closure f ist.env_bkt args + Tac2ffi.to_closure f args | GTacLet (false, el, e) -> let fold accu (na, e) = interp ist e >>= fun e -> @@ -82,7 +107,7 @@ let rec interp (ist : environment) = function let fixs = List.map map el in let fold accu (na, _, cls) = match na with | Anonymous -> accu - | Name id -> { ist with env_ist = Id.Map.add id cls accu.env_ist } + | Name id -> { env_ist = Id.Map.add id cls accu.env_ist } in let ist = List.fold_left fold ist fixs in (** Hack to make a cycle imperatively in the environment *) @@ -108,25 +133,24 @@ let rec interp (ist : environment) = function return (ValOpn (kn, Array.of_list el)) | GTacPrm (ml, el) -> Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> - Tac2env.interp_primitive ml (FrPrim ml :: ist.env_bkt) el + with_frame (FrPrim ml) (Tac2env.interp_primitive ml el) | GTacExt (tag, e) -> let tpe = Tac2env.interp_ml_object tag in - let ist = { ist with env_bkt = FrExtn (tag, e) :: ist.env_bkt } in - tpe.Tac2env.ml_interp ist e + with_frame (FrExtn (tag, e)) (tpe.Tac2env.ml_interp ist e) -and interp_app f = (); fun bt args -> match f with +and interp_app f = (); fun args -> match f with | { clos_env = ist; clos_var = ids; clos_exp = e; clos_ref = kn } -> let rec push ist ids args = match ids, args with - | [], [] -> interp ist e + | [], [] -> with_frame (FrLtac kn) (interp ist e) | [], _ :: _ -> - interp ist e >>= fun f -> Tac2ffi.to_closure f bt args + with_frame (FrLtac kn) (interp ist e) >>= fun f -> Tac2ffi.to_closure f args | _ :: _, [] -> let cls = { clos_ref = kn; clos_env = ist.env_ist; clos_var = ids; clos_exp = e } in let f = interp_app cls in return (ValCls f) | id :: ids, arg :: args -> push (push_name ist id arg) ids args in - push { env_ist = ist; env_bkt = FrLtac kn :: bt } ids args + push { env_ist = ist } ids args and interp_case ist e cse0 cse1 = match e with | ValInt n -> interp ist cse0.(n) diff --git a/src/tac2interp.mli b/src/tac2interp.mli index 48af59b178..1003e9f1eb 100644 --- a/src/tac2interp.mli +++ b/src/tac2interp.mli @@ -24,3 +24,11 @@ val set_env : environment -> Glob_term.unbound_ltac_var_map -> Glob_term.unbound exception LtacError of KerName.t * valexpr array * backtrace (** Ltac2-defined exceptions seen from OCaml side *) + +(** {5 Backtrace} *) + +val get_backtrace : backtrace Proofview.tactic + +val with_frame : frame -> 'a Proofview.tactic -> 'a Proofview.tactic + +val print_ltac2_backtrace : bool ref diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 14ad8695ca..28bcd6a1cf 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -18,7 +18,7 @@ module Value = Tac2ffi let return x = Proofview.tclUNIT x let v_unit = Value.of_unit () -let thaw bt f = (Value.to_closure f) bt [v_unit] +let thaw f = Value.to_closure f [v_unit] let to_pair f g = function | ValBlk (0, [| x; y |]) -> (f x, g y) @@ -126,8 +126,7 @@ and to_intro_patterns il = let to_destruction_arg = function | ValBlk (0, [| c |]) -> - (** FIXME: lost backtrace *) - let c = thaw [] c >>= fun c -> return (to_constr_with_bindings c) in + let c = thaw c >>= fun c -> return (to_constr_with_bindings c) in None, ElimOnConstr c | ValBlk (1, [| id |]) -> None, ElimOnIdent (Loc.tag (Value.to_ident id)) | ValBlk (2, [| n |]) -> None, ElimOnAnonHyp (Value.to_int n) @@ -155,7 +154,7 @@ let to_rewriting = function let orient = Value.to_option Value.to_bool orient in let repeat = to_multi repeat in (** FIXME: lost backtrace *) - let c = thaw [] c >>= fun c -> return (to_constr_with_bindings c) in + let c = thaw c >>= fun c -> return (to_constr_with_bindings c) in (orient, repeat, c) | _ -> assert false @@ -190,59 +189,59 @@ let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } let lift tac = tac <*> return v_unit let define_prim0 name tac = - let tac bt arg = match arg with + let tac arg = match arg with | [_] -> lift tac | _ -> assert false in Tac2env.define_primitive (pname name) tac let define_prim1 name tac = - let tac bt arg = match arg with - | [x] -> lift (tac bt x) + let tac arg = match arg with + | [x] -> lift (tac x) | _ -> assert false in Tac2env.define_primitive (pname name) tac let define_prim2 name tac = - let tac bt arg = match arg with - | [x; y] -> lift (tac bt x y) + let tac arg = match arg with + | [x; y] -> lift (tac x y) | _ -> assert false in Tac2env.define_primitive (pname name) tac let define_prim3 name tac = - let tac bt arg = match arg with - | [x; y; z] -> lift (tac bt x y z) + let tac arg = match arg with + | [x; y; z] -> lift (tac x y z) | _ -> assert false in Tac2env.define_primitive (pname name) tac let define_prim4 name tac = - let tac bt arg = match arg with - | [x; y; z; u] -> lift (tac bt x y z u) + let tac arg = match arg with + | [x; y; z; u] -> lift (tac x y z u) | _ -> assert false in Tac2env.define_primitive (pname name) tac let define_prim5 name tac = - let tac bt arg = match arg with - | [x; y; z; u; v] -> lift (tac bt x y z u v) + let tac arg = match arg with + | [x; y; z; u; v] -> lift (tac x y z u v) | _ -> assert false in Tac2env.define_primitive (pname name) tac (** Tactics from Tacexpr *) -let () = define_prim2 "tac_intros" begin fun _ ev ipat -> +let () = define_prim2 "tac_intros" begin fun ev ipat -> let ev = Value.to_bool ev in let ipat = to_intro_patterns ipat in Tactics.intros_patterns ev ipat end -let () = define_prim4 "tac_apply" begin fun bt adv ev cb ipat -> +let () = define_prim4 "tac_apply" begin fun adv ev cb ipat -> let adv = Value.to_bool adv in let ev = Value.to_bool ev in - let map_cb c = thaw bt c >>= fun c -> return (to_constr_with_bindings c) in + let map_cb c = thaw c >>= fun c -> return (to_constr_with_bindings c) in let cb = Value.to_list map_cb cb in let map p = Value.to_option (fun p -> Loc.tag (to_intro_pattern p)) p in let map_ipat p = to_pair Value.to_ident map p in @@ -250,20 +249,20 @@ let () = define_prim4 "tac_apply" begin fun bt adv ev cb ipat -> Tac2tactics.apply adv ev cb ipat end -let () = define_prim3 "tac_elim" begin fun _ ev c copt -> +let () = define_prim3 "tac_elim" begin fun ev c copt -> let ev = Value.to_bool ev in let c = to_constr_with_bindings c in let copt = Value.to_option to_constr_with_bindings copt in Tactics.elim ev None c copt end -let () = define_prim2 "tac_case" begin fun _ ev c -> +let () = define_prim2 "tac_case" begin fun ev c -> let ev = Value.to_bool ev in let c = to_constr_with_bindings c in Tactics.general_case_analysis ev None c end -let () = define_prim1 "tac_generalize" begin fun _ cl -> +let () = define_prim1 "tac_generalize" begin fun cl -> let cast = function | ValBlk (0, [| c; occs; na |]) -> ((to_occurrences Value.to_int occs, Value.to_constr c), to_name na) @@ -273,113 +272,113 @@ let () = define_prim1 "tac_generalize" begin fun _ cl -> Tactics.new_generalize_gen cl end -let () = define_prim3 "tac_assert" begin fun bt c tac ipat -> +let () = define_prim3 "tac_assert" begin fun c tac ipat -> let c = Value.to_constr c in - let of_tac t = Proofview.tclIGNORE (thaw bt t) in + let of_tac t = Proofview.tclIGNORE (thaw t) in let tac = Value.to_option (fun t -> Value.to_option of_tac t) tac in let ipat = Value.to_option (fun ipat -> Loc.tag (to_intro_pattern ipat)) ipat in Tactics.forward true tac ipat c end -let () = define_prim3 "tac_enough" begin fun bt c tac ipat -> +let () = define_prim3 "tac_enough" begin fun c tac ipat -> let c = Value.to_constr c in - let of_tac t = Proofview.tclIGNORE (thaw bt t) in + let of_tac t = Proofview.tclIGNORE (thaw t) in let tac = Value.to_option (fun t -> Value.to_option of_tac t) tac in let ipat = Value.to_option (fun ipat -> Loc.tag (to_intro_pattern ipat)) ipat in Tactics.forward false tac ipat c end -let () = define_prim2 "tac_pose" begin fun _ idopt c -> +let () = define_prim2 "tac_pose" begin fun idopt c -> let na = to_name idopt in let c = Value.to_constr c in Tactics.letin_tac None na c None Locusops.nowhere end -let () = define_prim4 "tac_set" begin fun bt ev idopt c cl -> +let () = define_prim4 "tac_set" begin fun ev idopt c cl -> let ev = Value.to_bool ev in let na = to_name idopt in let cl = to_clause cl in Proofview.tclEVARMAP >>= fun sigma -> - thaw bt c >>= fun c -> + thaw c >>= fun c -> let c = Value.to_constr c in Tactics.letin_pat_tac ev None na (sigma, c) cl end -let () = define_prim3 "tac_destruct" begin fun _ ev ic using -> +let () = define_prim3 "tac_destruct" begin fun ev ic using -> let ev = Value.to_bool ev in let ic = Value.to_list to_induction_clause ic in let using = Value.to_option to_constr_with_bindings using in Tac2tactics.induction_destruct false ev ic using end -let () = define_prim3 "tac_induction" begin fun _ ev ic using -> +let () = define_prim3 "tac_induction" begin fun ev ic using -> let ev = Value.to_bool ev in let ic = Value.to_list to_induction_clause ic in let using = Value.to_option to_constr_with_bindings using in Tac2tactics.induction_destruct true ev ic using end -let () = define_prim1 "tac_red" begin fun _ cl -> +let () = define_prim1 "tac_red" begin fun cl -> let cl = to_clause cl in Tactics.reduce (Red false) cl end -let () = define_prim1 "tac_hnf" begin fun _ cl -> +let () = define_prim1 "tac_hnf" begin fun cl -> let cl = to_clause cl in Tactics.reduce Hnf cl end -let () = define_prim3 "tac_simpl" begin fun _ flags where cl -> +let () = define_prim3 "tac_simpl" begin fun flags where cl -> let flags = to_red_flag flags in let where = Value.to_option to_pattern_with_occs where in let cl = to_clause cl in Tac2tactics.simpl flags where cl end -let () = define_prim2 "tac_cbv" begin fun _ flags cl -> +let () = define_prim2 "tac_cbv" begin fun flags cl -> let flags = to_red_flag flags in let cl = to_clause cl in Tac2tactics.cbv flags cl end -let () = define_prim2 "tac_cbn" begin fun _ flags cl -> +let () = define_prim2 "tac_cbn" begin fun flags cl -> let flags = to_red_flag flags in let cl = to_clause cl in Tac2tactics.cbn flags cl end -let () = define_prim2 "tac_lazy" begin fun _ flags cl -> +let () = define_prim2 "tac_lazy" begin fun flags cl -> let flags = to_red_flag flags in let cl = to_clause cl in Tac2tactics.lazy_ flags cl end -let () = define_prim2 "tac_unfold" begin fun _ refs cl -> +let () = define_prim2 "tac_unfold" begin fun refs cl -> let map v = to_pair Value.to_reference (fun occ -> to_occurrences to_int_or_var occ) v in let refs = Value.to_list map refs in let cl = to_clause cl in Tac2tactics.unfold refs cl end -let () = define_prim2 "tac_fold" begin fun _ args cl -> +let () = define_prim2 "tac_fold" begin fun args cl -> let args = Value.to_list Value.to_constr args in let cl = to_clause cl in Tactics.reduce (Fold args) cl end -let () = define_prim2 "tac_pattern" begin fun _ where cl -> +let () = define_prim2 "tac_pattern" begin fun where cl -> let where = Value.to_list to_constr_with_occs where in let cl = to_clause cl in Tactics.reduce (Pattern where) cl end -let () = define_prim2 "tac_vm" begin fun _ where cl -> +let () = define_prim2 "tac_vm" begin fun where cl -> let where = Value.to_option to_pattern_with_occs where in let cl = to_clause cl in Tac2tactics.vm where cl end -let () = define_prim2 "tac_native" begin fun _ where cl -> +let () = define_prim2 "tac_native" begin fun where cl -> let where = Value.to_option to_pattern_with_occs where in let cl = to_clause cl in Tac2tactics.native where cl @@ -388,102 +387,102 @@ end (** Reduction functions *) let define_red1 name tac = - let tac bt arg = match arg with - | [x] -> tac bt x >>= fun c -> Proofview.tclUNIT (Value.of_constr c) + let tac arg = match arg with + | [x] -> tac x >>= fun c -> Proofview.tclUNIT (Value.of_constr c) | _ -> assert false in Tac2env.define_primitive (pname name) tac let define_red2 name tac = - let tac bt arg = match arg with - | [x; y] -> tac bt x y >>= fun c -> Proofview.tclUNIT (Value.of_constr c) + let tac arg = match arg with + | [x; y] -> tac x y >>= fun c -> Proofview.tclUNIT (Value.of_constr c) | _ -> assert false in Tac2env.define_primitive (pname name) tac let define_red3 name tac = - let tac bt arg = match arg with - | [x; y; z] -> tac bt x y z >>= fun c -> Proofview.tclUNIT (Value.of_constr c) + let tac arg = match arg with + | [x; y; z] -> tac x y z >>= fun c -> Proofview.tclUNIT (Value.of_constr c) | _ -> assert false in Tac2env.define_primitive (pname name) tac -let () = define_red1 "eval_red" begin fun bt c -> +let () = define_red1 "eval_red" begin fun c -> let c = Value.to_constr c in - Tac2tactics.eval_red bt c + Tac2tactics.eval_red c end -let () = define_red1 "eval_hnf" begin fun bt c -> +let () = define_red1 "eval_hnf" begin fun c -> let c = Value.to_constr c in - Tac2tactics.eval_hnf bt c + Tac2tactics.eval_hnf c end -let () = define_red3 "eval_simpl" begin fun bt flags where c -> +let () = define_red3 "eval_simpl" begin fun flags where c -> let flags = to_red_flag flags in let where = Value.to_option to_pattern_with_occs where in let c = Value.to_constr c in - Tac2tactics.eval_simpl bt flags where c + Tac2tactics.eval_simpl flags where c end -let () = define_red2 "eval_cbv" begin fun bt flags c -> +let () = define_red2 "eval_cbv" begin fun flags c -> let flags = to_red_flag flags in let c = Value.to_constr c in - Tac2tactics.eval_cbv bt flags c + Tac2tactics.eval_cbv flags c end -let () = define_red2 "eval_cbn" begin fun bt flags c -> +let () = define_red2 "eval_cbn" begin fun flags c -> let flags = to_red_flag flags in let c = Value.to_constr c in - Tac2tactics.eval_cbn bt flags c + Tac2tactics.eval_cbn flags c end -let () = define_red2 "eval_lazy" begin fun bt flags c -> +let () = define_red2 "eval_lazy" begin fun flags c -> let flags = to_red_flag flags in let c = Value.to_constr c in - Tac2tactics.eval_lazy bt flags c + Tac2tactics.eval_lazy flags c end -let () = define_red2 "eval_unfold" begin fun bt refs c -> +let () = define_red2 "eval_unfold" begin fun refs c -> let map v = to_pair Value.to_reference (fun occ -> to_occurrences to_int_or_var occ) v in let refs = Value.to_list map refs in let c = Value.to_constr c in - Tac2tactics.eval_unfold bt refs c + Tac2tactics.eval_unfold refs c end -let () = define_red2 "eval_fold" begin fun bt args c -> +let () = define_red2 "eval_fold" begin fun args c -> let args = Value.to_list Value.to_constr args in let c = Value.to_constr c in - Tac2tactics.eval_fold bt args c + Tac2tactics.eval_fold args c end -let () = define_red2 "eval_pattern" begin fun bt where c -> +let () = define_red2 "eval_pattern" begin fun where c -> let where = Value.to_list (fun p -> to_pair Value.to_constr (fun occ -> to_occurrences to_int_or_var occ) p) where in let c = Value.to_constr c in - Tac2tactics.eval_pattern bt where c + Tac2tactics.eval_pattern where c end -let () = define_red2 "eval_vm" begin fun bt where c -> +let () = define_red2 "eval_vm" begin fun where c -> let where = Value.to_option to_pattern_with_occs where in let c = Value.to_constr c in - Tac2tactics.eval_vm bt where c + Tac2tactics.eval_vm where c end -let () = define_red2 "eval_native" begin fun bt where c -> +let () = define_red2 "eval_native" begin fun where c -> let where = Value.to_option to_pattern_with_occs where in let c = Value.to_constr c in - Tac2tactics.eval_native bt where c + Tac2tactics.eval_native where c end -let () = define_prim4 "tac_rewrite" begin fun bt ev rw cl by -> +let () = define_prim4 "tac_rewrite" begin fun ev rw cl by -> let ev = Value.to_bool ev in let rw = Value.to_list to_rewriting rw in let cl = to_clause cl in - let to_tac t = Proofview.tclIGNORE (thaw bt t) in + let to_tac t = Proofview.tclIGNORE (thaw t) in let by = Value.to_option to_tac by in Tac2tactics.rewrite ev rw cl by end -let () = define_prim4 "tac_inversion" begin fun bt knd arg pat ids -> +let () = define_prim4 "tac_inversion" begin fun knd arg pat ids -> let knd = to_inversion_kind knd in let arg = to_destruction_arg arg in let pat = Value.to_option (fun ipat -> Loc.tag (to_intro_pattern ipat)) pat in @@ -495,13 +494,13 @@ end let () = define_prim0 "tac_reflexivity" Tactics.intros_reflexivity -let () = define_prim2 "tac_move" begin fun _ id mv -> +let () = define_prim2 "tac_move" begin fun id mv -> let id = Value.to_ident id in let mv = to_move_location mv in Tactics.move_hyp id mv end -let () = define_prim2 "tac_intro" begin fun _ id mv -> +let () = define_prim2 "tac_intro" begin fun id mv -> let id = Value.to_option Value.to_ident id in let mv = Value.to_option to_move_location mv in let mv = Option.default MoveLast mv in @@ -518,69 +517,69 @@ END let () = define_prim0 "tac_assumption" Tactics.assumption -let () = define_prim1 "tac_transitivity" begin fun _ c -> +let () = define_prim1 "tac_transitivity" begin fun c -> let c = Value.to_constr c in Tactics.intros_transitivity (Some c) end let () = define_prim0 "tac_etransitivity" (Tactics.intros_transitivity None) -let () = define_prim1 "tac_cut" begin fun _ c -> +let () = define_prim1 "tac_cut" begin fun c -> let c = Value.to_constr c in Tactics.cut c end -let () = define_prim2 "tac_left" begin fun _ ev bnd -> +let () = define_prim2 "tac_left" begin fun ev bnd -> let ev = Value.to_bool ev in let bnd = to_bindings bnd in Tactics.left_with_bindings ev bnd end -let () = define_prim2 "tac_right" begin fun _ ev bnd -> +let () = define_prim2 "tac_right" begin fun ev bnd -> let ev = Value.to_bool ev in let bnd = to_bindings bnd in Tactics.right_with_bindings ev bnd end -let () = define_prim1 "tac_introsuntil" begin fun _ h -> +let () = define_prim1 "tac_introsuntil" begin fun h -> Tactics.intros_until (to_qhyp h) end -let () = define_prim1 "tac_exactnocheck" begin fun _ c -> +let () = define_prim1 "tac_exactnocheck" begin fun c -> Tactics.exact_no_check (Value.to_constr c) end -let () = define_prim1 "tac_vmcastnocheck" begin fun _ c -> +let () = define_prim1 "tac_vmcastnocheck" begin fun c -> Tactics.vm_cast_no_check (Value.to_constr c) end -let () = define_prim1 "tac_nativecastnocheck" begin fun _ c -> +let () = define_prim1 "tac_nativecastnocheck" begin fun c -> Tactics.native_cast_no_check (Value.to_constr c) end -let () = define_prim1 "tac_constructor" begin fun _ ev -> +let () = define_prim1 "tac_constructor" begin fun ev -> let ev = Value.to_bool ev in Tactics.any_constructor ev None end -let () = define_prim3 "tac_constructorn" begin fun _ ev n bnd -> +let () = define_prim3 "tac_constructorn" begin fun ev n bnd -> let ev = Value.to_bool ev in let n = Value.to_int n in let bnd = to_bindings bnd in Tactics.constructor_tac ev None n bnd end -let () = define_prim1 "tac_symmetry" begin fun _ cl -> +let () = define_prim1 "tac_symmetry" begin fun cl -> let cl = to_clause cl in Tactics.intros_symmetry cl end -let () = define_prim2 "tac_split" begin fun _ ev bnd -> +let () = define_prim2 "tac_split" begin fun ev bnd -> let ev = Value.to_bool ev in let bnd = to_bindings bnd in Tactics.split_with_bindings ev [bnd] end -let () = define_prim1 "tac_rename" begin fun _ ids -> +let () = define_prim1 "tac_rename" begin fun ids -> let map c = match Value.to_tuple c with | [|x; y|] -> (Value.to_ident x, Value.to_ident y) | _ -> assert false @@ -589,72 +588,72 @@ let () = define_prim1 "tac_rename" begin fun _ ids -> Tactics.rename_hyp ids end -let () = define_prim1 "tac_revert" begin fun _ ids -> +let () = define_prim1 "tac_revert" begin fun ids -> let ids = Value.to_list Value.to_ident ids in Tactics.revert ids end let () = define_prim0 "tac_admit" Proofview.give_up -let () = define_prim2 "tac_fix" begin fun _ idopt n -> +let () = define_prim2 "tac_fix" begin fun idopt n -> let idopt = Value.to_option Value.to_ident idopt in let n = Value.to_int n in Tactics.fix idopt n end -let () = define_prim1 "tac_cofix" begin fun _ idopt -> +let () = define_prim1 "tac_cofix" begin fun idopt -> let idopt = Value.to_option Value.to_ident idopt in Tactics.cofix idopt end -let () = define_prim1 "tac_clear" begin fun _ ids -> +let () = define_prim1 "tac_clear" begin fun ids -> let ids = Value.to_list Value.to_ident ids in Tactics.clear ids end -let () = define_prim1 "tac_keep" begin fun _ ids -> +let () = define_prim1 "tac_keep" begin fun ids -> let ids = Value.to_list Value.to_ident ids in Tactics.keep ids end -let () = define_prim1 "tac_clearbody" begin fun _ ids -> +let () = define_prim1 "tac_clearbody" begin fun ids -> let ids = Value.to_list Value.to_ident ids in Tactics.clear_body ids end (** Tactics from extratactics *) -let () = define_prim2 "tac_discriminate" begin fun _ ev arg -> +let () = define_prim2 "tac_discriminate" begin fun ev arg -> let ev = Value.to_bool ev in let arg = Value.to_option to_destruction_arg arg in Tac2tactics.discriminate ev arg end -let () = define_prim3 "tac_injection" begin fun _ ev ipat arg -> +let () = define_prim3 "tac_injection" begin fun ev ipat arg -> let ev = Value.to_bool ev in let ipat = Value.to_option to_intro_patterns ipat in let arg = Value.to_option to_destruction_arg arg in Tac2tactics.injection ev ipat arg end -let () = define_prim1 "tac_absurd" begin fun _ c -> +let () = define_prim1 "tac_absurd" begin fun c -> Contradiction.absurd (Value.to_constr c) end -let () = define_prim1 "tac_contradiction" begin fun _ c -> +let () = define_prim1 "tac_contradiction" begin fun c -> let c = Value.to_option to_constr_with_bindings c in Contradiction.contradiction c end -let () = define_prim4 "tac_autorewrite" begin fun bt all by ids cl -> +let () = define_prim4 "tac_autorewrite" begin fun all by ids cl -> let all = Value.to_bool all in - let by = Value.to_option (fun tac -> Proofview.tclIGNORE (thaw bt tac)) by in + let by = Value.to_option (fun tac -> Proofview.tclIGNORE (thaw tac)) by in let ids = Value.to_list Value.to_ident ids in let cl = to_clause cl in Tac2tactics.autorewrite ~all by ids cl end -let () = define_prim1 "tac_subst" begin fun _ ids -> +let () = define_prim1 "tac_subst" begin fun ids -> let ids = Value.to_list Value.to_ident ids in Equality.subst ids end @@ -663,43 +662,43 @@ let () = define_prim0 "tac_substall" (return () >>= fun () -> Equality.subst_all (** Auto *) -let () = define_prim3 "tac_trivial" begin fun bt dbg lems dbs -> +let () = define_prim3 "tac_trivial" begin fun dbg lems dbs -> let dbg = to_debug dbg in - let map c = thaw bt c >>= fun c -> return (Value.to_constr c) in + let map c = thaw c >>= fun c -> return (Value.to_constr c) in let lems = Value.to_list map lems in let dbs = Value.to_option (fun l -> Value.to_list Value.to_ident l) dbs in Tac2tactics.trivial dbg lems dbs end -let () = define_prim5 "tac_eauto" begin fun bt dbg n p lems dbs -> +let () = define_prim5 "tac_eauto" begin fun dbg n p lems dbs -> let dbg = to_debug dbg in let n = Value.to_option Value.to_int n in let p = Value.to_option Value.to_int p in - let map c = thaw bt c >>= fun c -> return (Value.to_constr c) in + let map c = thaw c >>= fun c -> return (Value.to_constr c) in let lems = Value.to_list map lems in let dbs = Value.to_option (fun l -> Value.to_list Value.to_ident l) dbs in Tac2tactics.eauto dbg n p lems dbs end -let () = define_prim4 "tac_auto" begin fun bt dbg n lems dbs -> +let () = define_prim4 "tac_auto" begin fun dbg n lems dbs -> let dbg = to_debug dbg in let n = Value.to_option Value.to_int n in - let map c = thaw bt c >>= fun c -> return (Value.to_constr c) in + let map c = thaw c >>= fun c -> return (Value.to_constr c) in let lems = Value.to_list map lems in let dbs = Value.to_option (fun l -> Value.to_list Value.to_ident l) dbs in Tac2tactics.auto dbg n lems dbs end -let () = define_prim4 "tac_newauto" begin fun bt dbg n lems dbs -> +let () = define_prim4 "tac_newauto" begin fun dbg n lems dbs -> let dbg = to_debug dbg in let n = Value.to_option Value.to_int n in - let map c = thaw bt c >>= fun c -> return (Value.to_constr c) in + let map c = thaw c >>= fun c -> return (Value.to_constr c) in let lems = Value.to_list map lems in let dbs = Value.to_option (fun l -> Value.to_list Value.to_ident l) dbs in Tac2tactics.new_auto dbg n lems dbs end -let () = define_prim3 "tac_typeclasses_eauto" begin fun bt str n dbs -> +let () = define_prim3 "tac_typeclasses_eauto" begin fun str n dbs -> let str = Value.to_option to_strategy str in let n = Value.to_option Value.to_int n in let dbs = Value.to_option (fun l -> Value.to_list Value.to_ident l) dbs in @@ -708,8 +707,8 @@ end (** Firstorder *) -let () = define_prim3 "tac_firstorder" begin fun bt tac refs ids -> - let tac = Value.to_option (fun t -> Proofview.tclIGNORE (thaw bt t)) tac in +let () = define_prim3 "tac_firstorder" begin fun tac refs ids -> + let tac = Value.to_option (fun t -> Proofview.tclIGNORE (thaw t)) tac in let refs = Value.to_list Value.to_reference refs in let ids = Value.to_list Value.to_ident ids in Tac2tactics.firstorder tac refs ids diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index dacbb898d3..aa2ee4711a 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -122,62 +122,62 @@ let native where cl = let where = Option.map map_pattern_with_occs where in Tactics.reduce (CbvNative where) cl -let eval_fun bt red c = - Tac2core.pf_apply bt begin fun env sigma -> +let eval_fun red c = + Tac2core.pf_apply begin fun env sigma -> let (redfun, _) = Redexpr.reduction_of_red_expr env red in let (sigma, ans) = redfun env sigma c in Proofview.Unsafe.tclEVARS sigma >>= fun () -> Proofview.tclUNIT ans end -let eval_red bt c = - eval_fun bt (Red false) c +let eval_red c = + eval_fun (Red false) c -let eval_hnf bt c = - eval_fun bt Hnf c +let eval_hnf c = + eval_fun Hnf c -let eval_simpl bt flags where c = +let eval_simpl flags where c = let where = Option.map map_pattern_with_occs where in Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> let flags = { flags with rConst } in - eval_fun bt (Simpl (flags, where)) c + eval_fun (Simpl (flags, where)) c -let eval_cbv bt flags c = +let eval_cbv flags c = Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> let flags = { flags with rConst } in - eval_fun bt (Cbv flags) c + eval_fun (Cbv flags) c -let eval_cbn bt flags c = +let eval_cbn flags c = Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> let flags = { flags with rConst } in - eval_fun bt (Cbn flags) c + eval_fun (Cbn flags) c -let eval_lazy bt flags c = +let eval_lazy flags c = Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> let flags = { flags with rConst } in - eval_fun bt (Lazy flags) c + eval_fun (Lazy flags) c -let eval_unfold bt occs c = +let eval_unfold occs c = let map (gr, occ) = get_evaluable_reference gr >>= fun gr -> Proofview.tclUNIT (occ, gr) in Proofview.Monad.List.map map occs >>= fun occs -> - eval_fun bt (Unfold occs) c + eval_fun (Unfold occs) c -let eval_fold bt cl c = - eval_fun bt (Fold cl) c +let eval_fold cl c = + eval_fun (Fold cl) c -let eval_pattern bt where c = +let eval_pattern where c = let where = List.map (fun (pat, occ) -> (occ, pat)) where in - eval_fun bt (Pattern where) c + eval_fun (Pattern where) c -let eval_vm bt where c = +let eval_vm where c = let where = Option.map map_pattern_with_occs where in - eval_fun bt (CbvVm where) c + eval_fun (CbvVm where) c -let eval_native bt where c = +let eval_native where c = let where = Option.map map_pattern_with_occs where in - eval_fun bt (CbvNative where) c + eval_fun (CbvNative where) c let on_destruction_arg tac ev arg = Proofview.Goal.enter begin fun gl -> diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli index 841e8fe762..f6825d84aa 100644 --- a/src/tac2tactics.mli +++ b/src/tac2tactics.mli @@ -56,28 +56,28 @@ val vm : (Pattern.constr_pattern * occurrences_expr) option -> clause -> unit ta val native : (Pattern.constr_pattern * occurrences_expr) option -> clause -> unit tactic -val eval_red : backtrace -> constr -> constr tactic +val eval_red : constr -> constr tactic -val eval_hnf : backtrace -> constr -> constr tactic +val eval_hnf : constr -> constr tactic -val eval_simpl : backtrace -> global_reference glob_red_flag -> +val eval_simpl : global_reference glob_red_flag -> (Pattern.constr_pattern * occurrences_expr) option -> constr -> constr tactic -val eval_cbv : backtrace -> global_reference glob_red_flag -> constr -> constr tactic +val eval_cbv : global_reference glob_red_flag -> constr -> constr tactic -val eval_cbn : backtrace -> global_reference glob_red_flag -> constr -> constr tactic +val eval_cbn : global_reference glob_red_flag -> constr -> constr tactic -val eval_lazy : backtrace -> global_reference glob_red_flag -> constr -> constr tactic +val eval_lazy : global_reference glob_red_flag -> constr -> constr tactic -val eval_unfold : backtrace -> (global_reference * occurrences_expr) list -> constr -> constr tactic +val eval_unfold : (global_reference * occurrences_expr) list -> constr -> constr tactic -val eval_fold : backtrace -> constr list -> constr -> constr tactic +val eval_fold : constr list -> constr -> constr tactic -val eval_pattern : backtrace -> (EConstr.t * occurrences_expr) list -> constr -> constr tactic +val eval_pattern : (EConstr.t * occurrences_expr) list -> constr -> constr tactic -val eval_vm : backtrace -> (Pattern.constr_pattern * occurrences_expr) option -> constr -> constr tactic +val eval_vm : (Pattern.constr_pattern * occurrences_expr) option -> constr -> constr tactic -val eval_native : backtrace -> (Pattern.constr_pattern * occurrences_expr) option -> constr -> constr tactic +val eval_native : (Pattern.constr_pattern * occurrences_expr) option -> constr -> constr tactic val discriminate : evars_flag -> destruction_arg option -> unit tactic -- cgit v1.2.3 From d577fe086794fda2edb3b98c12606e24c9c92ea1 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 7 Sep 2017 16:06:04 +0200 Subject: Fix coq/ltac2#21: Backtraces should print Ltac2 closures. --- src/tac2entries.ml | 4 ++-- src/tac2expr.mli | 3 ++- src/tac2interp.ml | 8 ++++++-- 3 files changed, 10 insertions(+), 5 deletions(-) diff --git a/src/tac2entries.ml b/src/tac2entries.ml index eed7eb6def..d596a61152 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -757,8 +757,8 @@ let _ = Goptions.declare_bool_option { } let pr_frame = function -| FrLtac None -> str "Call " -| FrLtac (Some kn) -> +| FrAnon e -> str "Call " +| FrLtac kn -> str "Call " ++ Libnames.pr_qualid (Tac2env.shortest_qualid_of_ltac (TacConstant kn)) | FrPrim ml -> str "Prim <" ++ str ml.mltac_plugin ++ str ":" ++ str ml.mltac_tactic ++ str ">" diff --git a/src/tac2expr.mli b/src/tac2expr.mli index 77e2cfef0e..0b9923f7b9 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -181,7 +181,8 @@ type strexpr = type tag = int type frame = -| FrLtac of ltac_constant option +| FrLtac of ltac_constant +| FrAnon of glb_tacexpr | FrPrim of ml_tactic_name | FrExtn : ('a, 'b) Tac2dyn.Arg.tag * 'b -> frame diff --git a/src/tac2interp.ml b/src/tac2interp.ml index f37b4f8e9c..19efeb7669 100644 --- a/src/tac2interp.ml +++ b/src/tac2interp.ml @@ -140,10 +140,14 @@ let rec interp (ist : environment) = function and interp_app f = (); fun args -> match f with | { clos_env = ist; clos_var = ids; clos_exp = e; clos_ref = kn } -> + let frame = match kn with + | None -> FrAnon e + | Some kn -> FrLtac kn + in let rec push ist ids args = match ids, args with - | [], [] -> with_frame (FrLtac kn) (interp ist e) + | [], [] -> with_frame frame (interp ist e) | [], _ :: _ -> - with_frame (FrLtac kn) (interp ist e) >>= fun f -> Tac2ffi.to_closure f args + with_frame frame (interp ist e) >>= fun f -> Tac2ffi.to_closure f args | _ :: _, [] -> let cls = { clos_ref = kn; clos_env = ist.env_ist; clos_var = ids; clos_exp = e } in let f = interp_app cls in -- cgit v1.2.3 From 643832c053e0255dd356231f4e5887db0228c2cd Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 7 Sep 2017 16:28:55 +0200 Subject: Slightly better printing for anonymous closures. --- src/tac2entries.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/tac2entries.ml b/src/tac2entries.ml index d596a61152..7cffdc6590 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -757,7 +757,7 @@ let _ = Goptions.declare_bool_option { } let pr_frame = function -| FrAnon e -> str "Call " +| FrAnon e -> str "Call {" ++ pr_glbexpr e ++ str "}" | FrLtac kn -> str "Call " ++ Libnames.pr_qualid (Tac2env.shortest_qualid_of_ltac (TacConstant kn)) | FrPrim ml -> -- cgit v1.2.3 From bdd594093b4fb7e46a6cae0135b6630d75bae6f6 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 7 Sep 2017 23:45:26 +0200 Subject: Fix coq/ltac2#22: Argument to Tactic_failure should be printed. We implement a printer for toplevel values and use it for exceptions in particular. --- src/ltac2_plugin.mlpack | 2 +- src/tac2entries.ml | 7 ++- src/tac2print.ml | 123 ++++++++++++++++++++++++++++++++++++++++++++++-- src/tac2print.mli | 9 ++++ 4 files changed, 133 insertions(+), 8 deletions(-) diff --git a/src/ltac2_plugin.mlpack b/src/ltac2_plugin.mlpack index 00ba5bc58e..70f97b9d1e 100644 --- a/src/ltac2_plugin.mlpack +++ b/src/ltac2_plugin.mlpack @@ -1,8 +1,8 @@ Tac2dyn Tac2env +Tac2ffi Tac2print Tac2intern -Tac2ffi Tac2interp Tac2entries Tac2quote diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 7cffdc6590..9a2693379b 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -768,8 +768,11 @@ let pr_frame = function obj.Tac2env.ml_print (Global.env ()) arg let () = register_handler begin function -| Tac2interp.LtacError (kn, _, bt) -> - let c = Tac2print.pr_constructor kn in (** FIXME *) +| Tac2interp.LtacError (kn, args, bt) -> + let t_exn = KerName.make2 Tac2env.coq_prefix (Label.make "exn") in + let v = ValOpn (kn, args) in + let t = GTypRef (Other t_exn, []) in + let c = Tac2print.pr_valexpr (Global.env ()) Evd.empty v t in let bt = if !Tac2interp.print_ltac2_backtrace then fnl () ++ str "Backtrace:" ++ fnl () ++ prlist_with_sep fnl pr_frame bt diff --git a/src/tac2print.ml b/src/tac2print.ml index 75ad2082d4..fd1d759cce 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -112,14 +112,14 @@ let pr_name = function let find_constructor n empty def = let rec find n = function | [] -> assert false - | (id, []) :: rem -> + | (id, []) as ans :: rem -> if empty then - if Int.equal n 0 then id + if Int.equal n 0 then ans else find (pred n) rem else find n rem - | (id, _ :: _) :: rem -> + | (id, _ :: _) as ans :: rem -> if not empty then - if Int.equal n 0 then id + if Int.equal n 0 then ans else find (pred n) rem else find n rem in @@ -130,7 +130,7 @@ let pr_internal_constructor tpe n is_const = | (_, GTydAlg data) -> data | _ -> assert false in - let id = find_constructor n is_const data.galg_constructors in + let (id, _) = find_constructor n is_const data.galg_constructors in let kn = change_kn_label tpe id in pr_constructor kn @@ -302,3 +302,116 @@ let pr_glbexpr_gen lvl c = let pr_glbexpr c = pr_glbexpr_gen E5 c + +(** Toplevel printers *) + +let rec subst_type subst (t : 'a glb_typexpr) = match t with +| GTypVar id -> subst.(id) +| GTypArrow (t1, t2) -> GTypArrow (subst_type subst t1, subst_type subst t2) +| GTypRef (qid, args) -> + GTypRef (qid, List.map (fun t -> subst_type subst t) args) + +let unfold kn args = + let (nparams, def) = Tac2env.interp_type kn in + match def with + | GTydDef (Some def) -> + let args = Array.of_list args in + Some (subst_type args def) + | _ -> None + +let rec kind t = match t with +| GTypVar id -> GTypVar id +| GTypRef (Other kn, tl) -> + begin match unfold kn tl with + | None -> t + | Some t -> kind t + end +| GTypArrow _ | GTypRef (Tuple _, _) -> t + +type val_printer = + { val_printer : 'a. Environ.env -> Evd.evar_map -> valexpr -> 'a glb_typexpr list -> Pp.t } + +let printers = ref KNmap.empty + +let register_val_printer kn pr = + printers := KNmap.add kn pr !printers + +open Tac2ffi + +let rec pr_valexpr env sigma v t = match kind t with +| GTypVar _ -> str "" +| GTypRef (Other kn, params) -> + let pr = try Some (KNmap.find kn !printers) with Not_found -> None in + begin match pr with + | Some pr -> pr.val_printer env sigma v params + | None -> + let n, repr = Tac2env.interp_type kn in + match repr with + | GTydDef None -> str "" + | GTydDef (Some _) -> + (** Shouldn't happen thanks to kind *) + assert false + | GTydAlg alg -> + begin match v with + | ValInt n -> pr_internal_constructor kn n true + | ValBlk (n, args) -> + let (id, tpe) = find_constructor n false alg.galg_constructors in + let knc = change_kn_label kn id in + let args = pr_constrargs env sigma params args tpe in + hv 2 (pr_constructor knc ++ spc () ++ str "(" ++ args ++ str ")") + | _ -> str "" + end + | GTydRec rcd -> str "{}" + | GTydOpn -> + begin match v with + | ValOpn (knc, [||]) -> pr_constructor knc + | ValOpn (knc, args) -> + let data = Tac2env.interp_constructor knc in + let args = pr_constrargs env sigma params args data.Tac2env.cdata_args in + hv 2 (pr_constructor knc ++ spc () ++ str "(" ++ args ++ str ")") + | _ -> str "" + end + end +| GTypArrow _ -> str "" +| GTypRef (Tuple _, tl) -> + let blk = Array.to_list (block.r_to v) in + if List.length blk == List.length tl then + let prs = List.map2 (fun v t -> pr_valexpr env sigma v t) blk tl in + hv 2 (str "(" ++ prlist_with_sep pr_comma (fun p -> p) prs ++ str ")") + else + str "" + +and pr_constrargs env sigma params args tpe = + let subst = Array.of_list params in + let tpe = List.map (fun t -> subst_type subst t) tpe in + let args = Array.to_list args in + let args = List.combine args tpe in + prlist_with_sep pr_comma (fun (v, t) -> pr_valexpr env sigma v t) args + +let register_init n f = + let kn = KerName.make2 Tac2env.coq_prefix (Label.make n) in + register_val_printer kn { val_printer = fun _ _ v _ -> f v } + +let () = register_init "int" begin fun n -> + let n = to_int n in + Pp.int n +end + +let () = register_init "string" begin fun s -> + let s = to_string s in + Pp.quote (Pp.str s) +end + +let () = register_init "ident" begin fun id -> + let id = to_ident id in + Pp.str "@" ++ Id.print id +end + +let () = register_init "message" begin fun pp -> + str "message:(" ++ to_pp pp ++ str ")" +end + +let () = register_init "err" begin fun e -> + let (e, _) = to_ext val_exn e in + str "err:(" ++ CErrors.print_no_report e ++ str ")" +end diff --git a/src/tac2print.mli b/src/tac2print.mli index 737e813ed3..01abd1efb1 100644 --- a/src/tac2print.mli +++ b/src/tac2print.mli @@ -29,6 +29,15 @@ val pr_projection : ltac_projection -> Pp.t val pr_glbexpr_gen : exp_level -> glb_tacexpr -> Pp.t val pr_glbexpr : glb_tacexpr -> Pp.t +(** {5 Printing values}*) + +type val_printer = + { val_printer : 'a. Environ.env -> Evd.evar_map -> valexpr -> 'a glb_typexpr list -> Pp.t } + +val register_val_printer : type_constant -> val_printer -> unit + +val pr_valexpr : Environ.env -> Evd.evar_map -> valexpr -> 'a glb_typexpr -> Pp.t + (** {5 Utilities} *) val int_name : unit -> (int -> string) -- cgit v1.2.3 From 9820b2c72cbf2da61cf44456334b38683799fd58 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 8 Sep 2017 01:39:51 +0200 Subject: Fix coq/ltac2#24: There should be a way to turn an exn into a message. --- src/tac2core.ml | 9 +++++++++ theories/Message.v | 3 +++ 2 files changed, 12 insertions(+) diff --git a/src/tac2core.ml b/src/tac2core.ml index 4a35442b04..f969183dce 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -37,6 +37,7 @@ let t_constr = coq_core "constr" let t_pattern = coq_core "pattern" let t_ident = coq_core "ident" let t_option = coq_core "option" +let t_exn = coq_core "exn" let t_reference = std_core "reference" let c_nil = coq_core "[]" @@ -188,6 +189,14 @@ let () = define1 "message_of_ident" ident begin fun c -> return (Value.of_pp pp) end +let () = define1 "message_of_exn" valexpr begin fun v -> + Proofview.tclENV >>= fun env -> + Proofview.tclEVARMAP >>= fun sigma -> + let pp = Tac2print.pr_valexpr env sigma v (GTypRef (Other Core.t_exn, [])) in + return (Value.of_pp pp) +end + + let () = define2 "message_concat" pp pp begin fun m1 m2 -> return (Value.of_pp (Pp.app m1 m2)) end diff --git a/theories/Message.v b/theories/Message.v index 45f4b221db..7bffe0746b 100644 --- a/theories/Message.v +++ b/theories/Message.v @@ -19,4 +19,7 @@ Ltac2 @ external of_ident : ident -> message := "ltac2" "message_of_ident". Ltac2 @ external of_constr : constr -> message := "ltac2" "message_of_constr". (** Panics if there is more than one goal under focus. *) +Ltac2 @ external of_exn : exn -> message := "ltac2" "message_of_exn". +(** Panics if there is more than one goal under focus. *) + Ltac2 @ external concat : message -> message -> message := "ltac2" "message_concat". -- cgit v1.2.3 From 2b66bf0083fd85cf2fc983dbca75b848194f897f Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 8 Sep 2017 02:32:54 +0200 Subject: Fix coq/ltac2#27: ... is not a particularly helpful printing of an error message. --- src/tac2print.ml | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/src/tac2print.ml b/src/tac2print.ml index fd1d759cce..ee45b33706 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -412,6 +412,7 @@ let () = register_init "message" begin fun pp -> end let () = register_init "err" begin fun e -> - let (e, _) = to_ext val_exn e in + let e = to_ext val_exn e in + let (e, _) = ExplainErr.process_vernac_interp_error ~allow_uncaught:true e in str "err:(" ++ CErrors.print_no_report e ++ str ")" end -- cgit v1.2.3 From 555a7cf0ce4457ecfbf68cd12dd0e801728f6662 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 8 Sep 2017 19:02:51 +0200 Subject: Using a dedicated argument for tactic quotations. This prevents having to go through an expensive phase of goal-building, when we can simply type-check the term. --- src/g_ltac2.ml4 | 4 ++-- src/tac2core.ml | 17 +++++++++++++++++ src/tac2env.ml | 1 + src/tac2env.mli | 1 + src/tac2intern.ml | 25 +++++++++++++++++++++++++ 5 files changed, 46 insertions(+), 2 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 5c285010e9..7295f31181 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -717,8 +717,8 @@ GEXTEND Gram let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in CAst.make ~loc:!@loc (CHole (None, IntroAnonymous, Some arg)) | test_dollar_ident; "$"; id = Prim.ident -> - let tac = Tac2quote.of_exact_var ~loc:!@loc (Loc.tag ~loc:!@loc id) in - let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in + let id = Loc.tag ~loc:!@loc id in + let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2_quotation) id in CAst.make ~loc:!@loc (CHole (None, IntroAnonymous, Some arg)) ] ] ; diff --git a/src/tac2core.ml b/src/tac2core.ml index f969183dce..8b83e501f9 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -912,6 +912,23 @@ let () = in Pretyping.register_constr_interp0 wit_ltac2 interp +let () = + let interp ist env sigma concl id = + let ist = Tac2interp.get_env ist in + let c = Id.Map.find id ist.env_ist in + let c = Value.to_constr c in + let evdref = ref sigma in + let () = Typing.e_check env evdref c concl in + (c, !evdref) + in + Pretyping.register_constr_interp0 wit_ltac2_quotation interp + +let () = + let pr_raw id = mt () in + let pr_glb id = str "$" ++ Id.print id in + let pr_top _ = mt () in + Genprint.register_print0 wit_ltac2_quotation pr_raw pr_glb pr_top + (** Ltac2 in Ltac1 *) let () = diff --git a/src/tac2env.ml b/src/tac2env.ml index ef2b44afb9..831c6a3b42 100644 --- a/src/tac2env.ml +++ b/src/tac2env.ml @@ -276,6 +276,7 @@ let std_prefix = (** Generic arguments *) let wit_ltac2 = Genarg.make0 "ltac2:value" +let wit_ltac2_quotation = Genarg.make0 "ltac2:quotation" let is_constructor qid = let (_, id) = repr_qualid qid in diff --git a/src/tac2env.mli b/src/tac2env.mli index 49c9910a44..89e22f07d5 100644 --- a/src/tac2env.mli +++ b/src/tac2env.mli @@ -132,6 +132,7 @@ val std_prefix : ModPath.t (** {5 Generic arguments} *) val wit_ltac2 : (raw_tacexpr, glb_tacexpr, Util.Empty.t) genarg_type +val wit_ltac2_quotation : (Id.t Loc.located, Id.t, Util.Empty.t) genarg_type (** {5 Helper functions} *) diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 2dcd8b8da3..0efd9a3005 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -25,6 +25,7 @@ let t_int = coq_type "int" let t_string = coq_type "string" let t_array = coq_type "array" let t_list = coq_type "list" +let t_constr = coq_type "constr" let c_nil = GTacCst (Other t_list, 0, []) let c_cons e el = GTacCst (Other t_list, 0, [e; el]) @@ -1511,3 +1512,27 @@ let () = in Genintern.register_intern0 wit_ltac2 intern let () = Genintern.register_subst0 wit_ltac2 subst_expr + +let () = + let open Genintern in + let intern ist (loc, id) = + let env = match Genintern.Store.get ist.extra ltac2_env with + | None -> + (** Only happens when Ltac2 is called from a constr or ltac1 quotation *) + let env = empty_env () in + if !Ltac_plugin.Tacintern.strict_check then env + else { env with env_str = false } + | Some env -> env + in + let t = + try Id.Map.find id env.env_var + with Not_found -> + CErrors.user_err ?loc (str "Unbound value " ++ Id.print id) + in + let t = fresh_mix_type_scheme env t in + let () = unify ?loc env t (GTypRef (Other t_constr, [])) in + (ist, id) + in + Genintern.register_intern0 wit_ltac2_quotation intern + +let () = Genintern.register_subst0 wit_ltac2_quotation (fun _ id -> id) -- cgit v1.2.3 From 0ece31492e4cf2813025aab3c4972bb769a3dea2 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 9 Sep 2017 02:02:52 +0200 Subject: Fix coq/ltac2#25: Control.case should not be able to catch Control.throw. When crossing constr boundaries, we mark exceptions as being fatal not to catch them. --- src/tac2core.ml | 18 +++++++++++++++--- 1 file changed, 15 insertions(+), 3 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 8b83e501f9..e01b9f3086 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -103,9 +103,12 @@ let err_matchfailure bt = let thaw f = f [v_unit] +let fatal_flag : unit Exninfo.t = Exninfo.make () +let fatal_info = Exninfo.add Exninfo.null fatal_flag () + let throw e = Tac2interp.get_backtrace >>= fun bt -> - Proofview.tclLIFT (Proofview.NonLogical.raise (e bt)) + Proofview.tclLIFT (Proofview.NonLogical.raise ~info:fatal_info (e bt)) let fail e = Tac2interp.get_backtrace >>= fun bt -> Proofview.tclZERO (e bt) @@ -571,6 +574,7 @@ end let () = define1 "throw" exn begin fun (e, info) -> Tac2interp.get_backtrace >>= fun bt -> let e = set_bt bt e in + let info = Exninfo.add info fatal_flag () in Proofview.tclLIFT (Proofview.NonLogical.raise ~info e) end @@ -784,16 +788,24 @@ let intern_constr self ist c = let (_, (c, _)) = Genintern.intern Stdarg.wit_constr ist c in (GlbVal c, gtypref t_constr) +let catchable_exception = function + | Logic_monad.Exception _ -> false + | e -> CErrors.noncritical e + let interp_constr flags ist c = let open Pretyping in let ist = to_lvar ist in pf_apply begin fun env sigma -> - Proofview.V82.wrap_exceptions begin fun () -> + try let (sigma, c) = understand_ltac flags env sigma ist WithoutTypeConstraint c in let c = ValExt (Value.val_constr, c) in Proofview.Unsafe.tclEVARS sigma >>= fun () -> Proofview.tclUNIT c - end + with e when catchable_exception e -> + let (e, info) = CErrors.push e in + match Exninfo.get info fatal_flag with + | None -> Proofview.tclZERO ~info e + | Some () -> Proofview.tclLIFT (Proofview.NonLogical.raise ~info e) end let () = -- cgit v1.2.3 From 254785a4a7f8373b5b0c4a289c2184cac3b5c420 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 9 Sep 2017 19:58:25 +0200 Subject: Moving Ltac2 backtraces to the Exninfo mechanism. I don't know why, but on CoqIDE this triggers a printing of the backtrace twice. This is not reproducible with coqtop. --- src/tac2core.ml | 46 +++++++++++++++++++++------------------------- src/tac2entries.ml | 23 +++++++++++++++-------- src/tac2entries.mli | 4 ++++ src/tac2ffi.ml | 6 +++--- src/tac2ffi.mli | 2 +- src/tac2interp.mli | 2 +- 6 files changed, 45 insertions(+), 38 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index e01b9f3086..ab87fa7739 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -87,17 +87,17 @@ let of_result f = function (** Stdlib exceptions *) -let err_notfocussed bt = - Tac2interp.LtacError (coq_core "Not_focussed", [||], bt) +let err_notfocussed = + Tac2interp.LtacError (coq_core "Not_focussed", [||]) -let err_outofbounds bt = - Tac2interp.LtacError (coq_core "Out_of_bounds", [||], bt) +let err_outofbounds = + Tac2interp.LtacError (coq_core "Out_of_bounds", [||]) -let err_notfound bt = - Tac2interp.LtacError (coq_core "Not_found", [||], bt) +let err_notfound = + Tac2interp.LtacError (coq_core "Not_found", [||]) -let err_matchfailure bt = - Tac2interp.LtacError (coq_core "Match_failure", [||], bt) +let err_matchfailure = + Tac2interp.LtacError (coq_core "Match_failure", [||]) (** Helper functions *) @@ -106,16 +106,18 @@ let thaw f = f [v_unit] let fatal_flag : unit Exninfo.t = Exninfo.make () let fatal_info = Exninfo.add Exninfo.null fatal_flag () -let throw e = +let set_bt info = Tac2interp.get_backtrace >>= fun bt -> - Proofview.tclLIFT (Proofview.NonLogical.raise ~info:fatal_info (e bt)) + Proofview.tclUNIT (Exninfo.add info Tac2entries.backtrace bt) -let fail e = - Tac2interp.get_backtrace >>= fun bt -> Proofview.tclZERO (e bt) +let throw ?(info = Exninfo.null) e = + set_bt info >>= fun info -> + let info = Exninfo.add info fatal_flag () in + Proofview.tclLIFT (Proofview.NonLogical.raise ~info e) -let set_bt bt e = match e with -| Tac2interp.LtacError (kn, args, _) -> Tac2interp.LtacError (kn, args, bt) -| e -> e +let fail ?(info = Exninfo.null) e = + set_bt info >>= fun info -> + Proofview.tclZERO ~info e let return x = Proofview.tclUNIT x let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } @@ -572,19 +574,14 @@ end (** Error *) let () = define1 "throw" exn begin fun (e, info) -> - Tac2interp.get_backtrace >>= fun bt -> - let e = set_bt bt e in - let info = Exninfo.add info fatal_flag () in - Proofview.tclLIFT (Proofview.NonLogical.raise ~info e) + throw ~info e end (** Control *) (** exn -> 'a *) let () = define1 "zero" exn begin fun (e, info) -> - Tac2interp.get_backtrace >>= fun bt -> - let e = set_bt bt e in - Proofview.tclZERO ~info e + fail ~info e end (** (unit -> 'a) -> (exn -> 'a) -> 'a *) @@ -624,8 +621,7 @@ let () = define1 "case" closure begin fun f -> let k = function | [e] -> let (e, info) = Value.to_exn e in - Tac2interp.get_backtrace >>= fun bt -> - let e = set_bt bt e in + set_bt info >>= fun info -> k (e, info) | _ -> assert false in @@ -805,7 +801,7 @@ let interp_constr flags ist c = let (e, info) = CErrors.push e in match Exninfo.get info fatal_flag with | None -> Proofview.tclZERO ~info e - | Some () -> Proofview.tclLIFT (Proofview.NonLogical.raise ~info e) + | Some () -> throw ~info e end let () = diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 9a2693379b..208231b814 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -756,6 +756,8 @@ let _ = Goptions.declare_bool_option { Goptions.optwrite = (fun b -> Tac2interp.print_ltac2_backtrace := b); } +let backtrace : backtrace Exninfo.t = Exninfo.make () + let pr_frame = function | FrAnon e -> str "Call {" ++ pr_glbexpr e ++ str "}" | FrLtac kn -> @@ -768,21 +770,26 @@ let pr_frame = function obj.Tac2env.ml_print (Global.env ()) arg let () = register_handler begin function -| Tac2interp.LtacError (kn, args, bt) -> +| Tac2interp.LtacError (kn, args) -> let t_exn = KerName.make2 Tac2env.coq_prefix (Label.make "exn") in let v = ValOpn (kn, args) in let t = GTypRef (Other t_exn, []) in let c = Tac2print.pr_valexpr (Global.env ()) Evd.empty v t in - let bt = - if !Tac2interp.print_ltac2_backtrace then - fnl () ++ str "Backtrace:" ++ fnl () ++ prlist_with_sep fnl pr_frame bt - else - mt () - in - hov 0 (str "Uncaught Ltac2 exception:" ++ spc () ++ hov 0 c) ++ bt + hov 0 (str "Uncaught Ltac2 exception:" ++ spc () ++ hov 0 c) | _ -> raise Unhandled end +let () = ExplainErr.register_additional_error_info begin fun (e, info) -> + if !Tac2interp.print_ltac2_backtrace then + let bt = Exninfo.get info backtrace in + let bt = Option.default [] bt in + let bt = + str "Backtrace:" ++ fnl () ++ prlist_with_sep fnl pr_frame bt ++ fnl () + in + Some (Loc.tag @@ Some bt) + else raise Exit +end + (** Printing *) let print_ltac ref = diff --git a/src/tac2entries.mli b/src/tac2entries.mli index dda1653593..7c71130402 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -49,6 +49,10 @@ val print_ltac : Libnames.reference -> unit (** Evaluate a tactic expression in the current environment *) val call : default:bool -> raw_tacexpr -> unit +(** {5 Toplevel exceptions} *) + +val backtrace : backtrace Exninfo.t + (** {5 Parsing entries} *) module Pltac : diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index fb97177c4d..16c17cce62 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -43,7 +43,7 @@ match Val.eq tag tag' with (** Exception *) -exception LtacError of KerName.t * valexpr array * backtrace +exception LtacError of KerName.t * valexpr array (** Conversion functions *) @@ -169,7 +169,7 @@ let internal_err = (** FIXME: handle backtrace in Ltac2 exceptions *) let of_exn c = match fst c with -| LtacError (kn, c, _) -> ValOpn (kn, c) +| LtacError (kn, c) -> ValOpn (kn, c) | _ -> ValOpn (internal_err, [|of_ext val_exn c|]) let to_exn c = match c with @@ -177,7 +177,7 @@ let to_exn c = match c with if Names.KerName.equal kn internal_err then to_ext val_exn c.(0) else - (LtacError (kn, c, []), Exninfo.null) + (LtacError (kn, c), Exninfo.null) | _ -> assert false let exn = { diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index dfc87f7db3..05f4d210ab 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -121,5 +121,5 @@ val val_exn : Exninfo.iexn Tac2dyn.Val.tag (** Exception *) -exception LtacError of KerName.t * valexpr array * backtrace +exception LtacError of KerName.t * valexpr array (** Ltac2-defined exceptions seen from OCaml side *) diff --git a/src/tac2interp.mli b/src/tac2interp.mli index 1003e9f1eb..3acca72367 100644 --- a/src/tac2interp.mli +++ b/src/tac2interp.mli @@ -22,7 +22,7 @@ val set_env : environment -> Glob_term.unbound_ltac_var_map -> Glob_term.unbound (** {5 Exceptions} *) -exception LtacError of KerName.t * valexpr array * backtrace +exception LtacError of KerName.t * valexpr array (** Ltac2-defined exceptions seen from OCaml side *) (** {5 Backtrace} *) -- cgit v1.2.3 From a059c181c3b2d6c1f9c3682c15270c0942430f39 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 9 Sep 2017 20:40:11 +0200 Subject: Fix coq/ltac2#18: Terms should show a backtrace when Set Ltac2 Backtrace is set. --- src/tac2core.ml | 1 + 1 file changed, 1 insertion(+) diff --git a/src/tac2core.ml b/src/tac2core.ml index ab87fa7739..9920b4c805 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -799,6 +799,7 @@ let interp_constr flags ist c = Proofview.tclUNIT c with e when catchable_exception e -> let (e, info) = CErrors.push e in + set_bt info >>= fun info -> match Exninfo.get info fatal_flag with | None -> Proofview.tclZERO ~info e | Some () -> throw ~info e -- cgit v1.2.3 From 91a9313fbe24dfb0c9b7fcaa31e3c11bf055450a Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 9 Sep 2017 20:43:32 +0200 Subject: Fix coq/ltac2#26: Ltac1 gives no backtraces. --- src/tac2core.ml | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 9920b4c805..f19f45fa4e 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -895,7 +895,9 @@ let () = let ist = Ltac_plugin.Tacinterp.default_ist () in (** FUCK YOU API *) let ist = { ist with API.Geninterp.lfun = (Obj.magic lfun) } in - (Obj.magic Ltac_plugin.Tacinterp.eval_tactic_ist ist tac : unit Proofview.tactic) >>= fun () -> + let tac = (Obj.magic Ltac_plugin.Tacinterp.eval_tactic_ist ist tac : unit Proofview.tactic) in + let wrap (e, info) = set_bt info >>= fun info -> Proofview.tclZERO ~info e in + Proofview.tclOR tac wrap >>= fun () -> return v_unit in let subst s tac = Genintern.substitute Ltac_plugin.Tacarg.wit_tactic s tac in -- cgit v1.2.3 From 6bc632021b47103b57abb34836766a57198d9cb4 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 9 Sep 2017 20:45:48 +0200 Subject: Update backtraces only when the Ltac2 Backtrace flag is set. --- src/tac2core.ml | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index f19f45fa4e..7bd0164b4d 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -107,8 +107,10 @@ let fatal_flag : unit Exninfo.t = Exninfo.make () let fatal_info = Exninfo.add Exninfo.null fatal_flag () let set_bt info = - Tac2interp.get_backtrace >>= fun bt -> - Proofview.tclUNIT (Exninfo.add info Tac2entries.backtrace bt) + if !Tac2interp.print_ltac2_backtrace then + Tac2interp.get_backtrace >>= fun bt -> + Proofview.tclUNIT (Exninfo.add info Tac2entries.backtrace bt) + else Proofview.tclUNIT info let throw ?(info = Exninfo.null) e = set_bt info >>= fun info -> -- cgit v1.2.3 From c7c1f9b2da838a604c479bb2bc162fef621524ed Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 9 Sep 2017 20:46:59 +0200 Subject: If backtrace is missing, don't print it. --- src/tac2entries.ml | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 208231b814..26f96f7d72 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -782,7 +782,10 @@ end let () = ExplainErr.register_additional_error_info begin fun (e, info) -> if !Tac2interp.print_ltac2_backtrace then let bt = Exninfo.get info backtrace in - let bt = Option.default [] bt in + let bt = match bt with + | Some bt -> bt + | None -> raise Exit + in let bt = str "Backtrace:" ++ fnl () ++ prlist_with_sep fnl pr_frame bt ++ fnl () in -- cgit v1.2.3 From 46095430ed07306ba3380ea8192540793c5c0a26 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 13 Sep 2017 21:44:43 +0200 Subject: Adding quotations for the generalize tactic. --- src/g_ltac2.ml4 | 18 +++++++++++++++++- src/tac2core.ml | 1 + src/tac2entries.ml | 1 + src/tac2entries.mli | 1 + src/tac2qexpr.mli | 5 +++++ src/tac2quote.ml | 10 ++++++++++ src/tac2quote.mli | 2 ++ theories/Notations.v | 6 ++++++ 8 files changed, 43 insertions(+), 1 deletion(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 7295f31181..dad1b71685 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -369,7 +369,7 @@ GEXTEND Gram GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause q_rewriting q_clause q_dispatch q_occurrences q_strategy_flag q_destruction_arg q_reference q_with_bindings q_constr_matching - q_hintdb q_move_location; + q_hintdb q_move_location q_generalizations; anti: [ [ "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] ; @@ -702,6 +702,22 @@ GEXTEND Gram q_move_location: [ [ mv = move_location -> mv ] ] ; + opt_generalization: + [ [ occs = occs; na = OPT ident_or_anti -> (occs, na) + | -> Loc.tag ~loc:!@loc QAllOccurrences, None + ] ] + ; + generalization: + [ [ c = Constr.constr; (occs, na) = opt_generalization -> + Loc.tag ~loc:!@loc (c, occs, na) + ] ] + ; + generalizations: + [ [ g = LIST1 generalization SEP "," -> Loc.tag ~loc:!@loc g ] ] + ; + q_generalizations: + [ [ g = generalizations -> g ] ] + ; END (** Extension of constr syntax *) diff --git a/src/tac2core.ml b/src/tac2core.ml index 7bd0164b4d..1eef098311 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -1118,6 +1118,7 @@ let () = add_expr_scope "dispatch" q_dispatch Tac2quote.of_dispatch let () = add_expr_scope "strategy" q_strategy_flag Tac2quote.of_strategy_flag let () = add_expr_scope "reference" q_reference Tac2quote.of_reference let () = add_expr_scope "move_location" q_move_location Tac2quote.of_move_location +let () = add_expr_scope "generalizations" q_generalizations Tac2quote.of_generalizations let () = add_expr_scope "constr_matching" q_constr_matching Tac2quote.of_constr_matching let () = add_generic_scope "constr" Pcoq.Constr.constr Tac2quote.wit_constr diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 26f96f7d72..2025753aec 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -40,6 +40,7 @@ let q_strategy_flag = Pcoq.Gram.entry_create "tactic:q_strategy_flag" let q_constr_matching = Pcoq.Gram.entry_create "tactic:q_constr_matching" let q_hintdb = Pcoq.Gram.entry_create "tactic:q_hintdb" let q_move_location = Pcoq.Gram.entry_create "tactic:q_move_location" +let q_generalizations = Pcoq.Gram.entry_create "tactic:q_generalizations" end (** Tactic definition *) diff --git a/src/tac2entries.mli b/src/tac2entries.mli index 7c71130402..d91032c2a8 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -79,6 +79,7 @@ val q_strategy_flag : strategy_flag Pcoq.Gram.entry val q_constr_matching : constr_matching Pcoq.Gram.entry val q_hintdb : hintdb Pcoq.Gram.entry val q_move_location : move_location Pcoq.Gram.entry +val q_generalizations : generalizations Pcoq.Gram.entry end (** {5 Hooks} *) diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index 7d02022e07..59c7ad6f3f 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -139,3 +139,8 @@ type move_location_r = | QMoveLast type move_location = move_location_r located + +type generalization = + (Constrexpr.constr_expr * occurrences * Id.t located or_anti option) located + +type generalizations = generalization list located diff --git a/src/tac2quote.ml b/src/tac2quote.ml index f14612d58f..a5c5c7d31f 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -370,3 +370,13 @@ let of_move_location (loc, mv) = match mv with | QMoveBefore id -> std_constructor ?loc "MoveBefore" [of_anti of_ident id] | QMoveFirst -> std_constructor ?loc "MoveFirst" [] | QMoveLast -> std_constructor ?loc "MoveLast" [] + +let of_generalization (loc, (c, occ, na)) = + of_tuple ?loc [ + of_open_constr c; + of_occurrences occ; + of_option (fun id -> of_anti of_ident id) na; + ] + +let of_generalizations (loc, l) = + of_list ?loc of_generalization l diff --git a/src/tac2quote.mli b/src/tac2quote.mli index 9f42c60042..d31ce455eb 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -71,6 +71,8 @@ val of_exact_hyp : ?loc:Loc.t -> Id.t located -> raw_tacexpr val of_exact_var : ?loc:Loc.t -> Id.t located -> raw_tacexpr (** id ↦ 'Control.refine (fun () => Control.hyp @id') *) +val of_generalizations : generalizations -> raw_tacexpr + val of_dispatch : dispatch -> raw_tacexpr val of_strategy_flag : strategy_flag -> raw_tacexpr diff --git a/theories/Notations.v b/theories/Notations.v index 9b39942ca5..e393002f55 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -284,6 +284,12 @@ Ltac2 Notation "einduction" use(thunk(opt(seq("using", constr, with_bindings)))) := induction0 true ic use. +Ltac2 generalize0 gen := + enter_h false (fun _ gen => Std.generalize gen) gen. + +Ltac2 Notation "generalize" gen(thunk(generalizations)) := + generalize0 gen. + Ltac2 destruct0 ev ic use := let f ev use := Std.destruct ev ic use in enter_h ev f use. -- cgit v1.2.3 From c6fb27a363b38c30c7f9953d5f74e9eb98a26387 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 13 Sep 2017 22:53:29 +0200 Subject: Better printing for list literals. --- src/tac2print.ml | 72 +++++++++++++++++++++++++++++++++++++++----------------- 1 file changed, 50 insertions(+), 22 deletions(-) diff --git a/src/tac2print.ml b/src/tac2print.ml index ee45b33706..7113b01610 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -21,6 +21,10 @@ let change_kn_label kn id = let paren p = hov 2 (str "(" ++ p ++ str ")") +let t_list = + KerName.make2 Tac2env.coq_prefix (Label.of_id (Id.of_string "list")) + + (** Type printing *) type typ_level = @@ -187,28 +191,7 @@ let pr_glbexpr_gen lvl c = in paren (prlist_with_sep (fun () -> str "," ++ spc ()) (pr_glbexpr E1) cl) | GTacCst (Other tpe, n, cl) -> - begin match Tac2env.interp_type tpe with - | _, GTydAlg def -> - let paren = match lvl with - | E0 -> paren - | E1 | E2 | E3 | E4 | E5 -> fun x -> x - in - let cstr = pr_internal_constructor tpe n (List.is_empty cl) in - let cl = match cl with - | [] -> mt () - | _ -> spc () ++ pr_sequence (pr_glbexpr E0) cl - in - paren (hov 2 (cstr ++ cl)) - | _, GTydRec def -> - let args = List.combine def cl in - let pr_arg ((id, _, _), arg) = - let kn = change_kn_label tpe id in - pr_projection kn ++ spc () ++ str ":=" ++ spc () ++ pr_glbexpr E1 arg - in - let args = prlist_with_sep (fun () -> str ";" ++ spc ()) pr_arg args in - hv 0 (str "{" ++ spc () ++ args ++ spc () ++ str "}") - | _, (GTydDef _ | GTydOpn) -> assert false - end + pr_applied_constructor lvl tpe n cl | GTacCse (e, info, cst_br, ncst_br) -> let e = pr_glbexpr E5 e in let br = match info with @@ -297,9 +280,54 @@ let pr_glbexpr_gen lvl c = in hov 0 (str "@external" ++ spc () ++ qstring prm.mltac_plugin ++ spc () ++ qstring prm.mltac_tactic ++ args) + and pr_applied_constructor lvl tpe n cl = + let _, data = Tac2env.interp_type tpe in + if KerName.equal tpe t_list then + let rec factorize accu = function + | GTacCst (_, 0, []) -> accu, None + | GTacCst (_, 0, [e; l]) -> factorize (e :: accu) l + | e -> accu, Some e + in + let l, e = factorize [] (GTacCst (Other tpe, n, cl)) in + match e with + | None -> + let pr e = pr_glbexpr E4 e in + hov 2 (str "[" ++ prlist_with_sep pr_semicolon pr (List.rev l) ++ str "]") + | Some e -> + let paren = match lvl with + | E0 | E1 | E2 -> paren + | E3 | E4 | E5 -> fun x -> x + in + let pr e = pr_glbexpr E1 e in + let pr_cons () = spc () ++ str "::" ++ spc () in + paren (hov 2 (prlist_with_sep pr_cons pr (List.rev (e :: l)))) + else match data with + | GTydAlg def -> + let paren = match lvl with + | E0 -> + if List.is_empty cl then fun x -> x else paren + | E1 | E2 | E3 | E4 | E5 -> fun x -> x + in + let cstr = pr_internal_constructor tpe n (List.is_empty cl) in + let cl = match cl with + | [] -> mt () + | _ -> spc () ++ pr_sequence (pr_glbexpr E0) cl + in + paren (hov 2 (cstr ++ cl)) + | GTydRec def -> + let args = List.combine def cl in + let pr_arg ((id, _, _), arg) = + let kn = change_kn_label tpe id in + pr_projection kn ++ spc () ++ str ":=" ++ spc () ++ pr_glbexpr E1 arg + in + let args = prlist_with_sep (fun () -> str ";" ++ spc ()) pr_arg args in + hv 0 (str "{" ++ spc () ++ args ++ spc () ++ str "}") + | (GTydDef _ | GTydOpn) -> assert false in hov 0 (pr_glbexpr lvl c) + + let pr_glbexpr c = pr_glbexpr_gen E5 c -- cgit v1.2.3 From 4ed40a9427f67ab6091f1af5457ffdec5e156d12 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 14 Sep 2017 00:02:39 +0200 Subject: Use a simpler syntax for generalize grammar. This removes the use for a quotation. --- src/g_ltac2.ml4 | 18 +----------------- src/tac2core.ml | 1 - src/tac2entries.ml | 1 - src/tac2entries.mli | 1 - src/tac2qexpr.mli | 5 ----- src/tac2quote.ml | 10 ---------- src/tac2quote.mli | 2 -- theories/Notations.v | 3 ++- 8 files changed, 3 insertions(+), 38 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index dad1b71685..7295f31181 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -369,7 +369,7 @@ GEXTEND Gram GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause q_rewriting q_clause q_dispatch q_occurrences q_strategy_flag q_destruction_arg q_reference q_with_bindings q_constr_matching - q_hintdb q_move_location q_generalizations; + q_hintdb q_move_location; anti: [ [ "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] ; @@ -702,22 +702,6 @@ GEXTEND Gram q_move_location: [ [ mv = move_location -> mv ] ] ; - opt_generalization: - [ [ occs = occs; na = OPT ident_or_anti -> (occs, na) - | -> Loc.tag ~loc:!@loc QAllOccurrences, None - ] ] - ; - generalization: - [ [ c = Constr.constr; (occs, na) = opt_generalization -> - Loc.tag ~loc:!@loc (c, occs, na) - ] ] - ; - generalizations: - [ [ g = LIST1 generalization SEP "," -> Loc.tag ~loc:!@loc g ] ] - ; - q_generalizations: - [ [ g = generalizations -> g ] ] - ; END (** Extension of constr syntax *) diff --git a/src/tac2core.ml b/src/tac2core.ml index 1eef098311..7bd0164b4d 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -1118,7 +1118,6 @@ let () = add_expr_scope "dispatch" q_dispatch Tac2quote.of_dispatch let () = add_expr_scope "strategy" q_strategy_flag Tac2quote.of_strategy_flag let () = add_expr_scope "reference" q_reference Tac2quote.of_reference let () = add_expr_scope "move_location" q_move_location Tac2quote.of_move_location -let () = add_expr_scope "generalizations" q_generalizations Tac2quote.of_generalizations let () = add_expr_scope "constr_matching" q_constr_matching Tac2quote.of_constr_matching let () = add_generic_scope "constr" Pcoq.Constr.constr Tac2quote.wit_constr diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 2025753aec..26f96f7d72 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -40,7 +40,6 @@ let q_strategy_flag = Pcoq.Gram.entry_create "tactic:q_strategy_flag" let q_constr_matching = Pcoq.Gram.entry_create "tactic:q_constr_matching" let q_hintdb = Pcoq.Gram.entry_create "tactic:q_hintdb" let q_move_location = Pcoq.Gram.entry_create "tactic:q_move_location" -let q_generalizations = Pcoq.Gram.entry_create "tactic:q_generalizations" end (** Tactic definition *) diff --git a/src/tac2entries.mli b/src/tac2entries.mli index d91032c2a8..7c71130402 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -79,7 +79,6 @@ val q_strategy_flag : strategy_flag Pcoq.Gram.entry val q_constr_matching : constr_matching Pcoq.Gram.entry val q_hintdb : hintdb Pcoq.Gram.entry val q_move_location : move_location Pcoq.Gram.entry -val q_generalizations : generalizations Pcoq.Gram.entry end (** {5 Hooks} *) diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index 59c7ad6f3f..7d02022e07 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -139,8 +139,3 @@ type move_location_r = | QMoveLast type move_location = move_location_r located - -type generalization = - (Constrexpr.constr_expr * occurrences * Id.t located or_anti option) located - -type generalizations = generalization list located diff --git a/src/tac2quote.ml b/src/tac2quote.ml index a5c5c7d31f..f14612d58f 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -370,13 +370,3 @@ let of_move_location (loc, mv) = match mv with | QMoveBefore id -> std_constructor ?loc "MoveBefore" [of_anti of_ident id] | QMoveFirst -> std_constructor ?loc "MoveFirst" [] | QMoveLast -> std_constructor ?loc "MoveLast" [] - -let of_generalization (loc, (c, occ, na)) = - of_tuple ?loc [ - of_open_constr c; - of_occurrences occ; - of_option (fun id -> of_anti of_ident id) na; - ] - -let of_generalizations (loc, l) = - of_list ?loc of_generalization l diff --git a/src/tac2quote.mli b/src/tac2quote.mli index d31ce455eb..9f42c60042 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -71,8 +71,6 @@ val of_exact_hyp : ?loc:Loc.t -> Id.t located -> raw_tacexpr val of_exact_var : ?loc:Loc.t -> Id.t located -> raw_tacexpr (** id ↦ 'Control.refine (fun () => Control.hyp @id') *) -val of_generalizations : generalizations -> raw_tacexpr - val of_dispatch : dispatch -> raw_tacexpr val of_strategy_flag : strategy_flag -> raw_tacexpr diff --git a/theories/Notations.v b/theories/Notations.v index e393002f55..9b2b04f9e6 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -287,7 +287,8 @@ Ltac2 Notation "einduction" Ltac2 generalize0 gen := enter_h false (fun _ gen => Std.generalize gen) gen. -Ltac2 Notation "generalize" gen(thunk(generalizations)) := +Ltac2 Notation "generalize" + gen(thunk(list1(seq (open_constr, occurrences, opt(seq("as", ident))), ","))) := generalize0 gen. Ltac2 destruct0 ev ic use := -- cgit v1.2.3 From c53fb4be8c65a89dd03d4aedc2fc65d9807da915 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 14 Sep 2017 00:15:09 +0200 Subject: Binding the pose/set family of tactics. --- src/g_ltac2.ml4 | 37 ++++++++++++++++++++++++++++++++++++- src/tac2core.ml | 1 + src/tac2entries.ml | 1 + src/tac2entries.mli | 1 + src/tac2qexpr.mli | 2 ++ src/tac2quote.ml | 3 +++ src/tac2quote.mli | 2 ++ src/tac2stdlib.ml | 7 +++---- tests/example2.v | 14 ++++++++++++++ theories/Notations.v | 27 +++++++++++++++++++++------ theories/Std.v | 2 +- 11 files changed, 85 insertions(+), 12 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 7295f31181..dfd586d5ef 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -40,6 +40,27 @@ let test_lpar_idnum_coloneq = | _ -> err ()) | _ -> err ()) +(* Hack to recognize "(x := t)" and "($x := t)" *) +let test_lpar_coloneq = + Gram.Entry.of_parser "test_coloneq" + (fun strm -> + match stream_nth 0 strm with + | KEYWORD "(" -> + (match stream_nth 1 strm with + | IDENT _ -> + (match stream_nth 2 strm with + | KEYWORD ":=" -> () + | _ -> err ()) + | KEYWORD "$" -> + (match stream_nth 2 strm with + | IDENT _ -> + (match stream_nth 3 strm with + | KEYWORD ":=" -> () + | _ -> err ()) + | _ -> err ()) + | _ -> err ()) + | _ -> err ()) + (* Hack to recognize "(x)" *) let test_lpar_id_rpar = Gram.Entry.of_parser "lpar_id_coloneq" @@ -369,7 +390,7 @@ GEXTEND Gram GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause q_rewriting q_clause q_dispatch q_occurrences q_strategy_flag q_destruction_arg q_reference q_with_bindings q_constr_matching - q_hintdb q_move_location; + q_hintdb q_move_location q_pose; anti: [ [ "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] ; @@ -702,6 +723,20 @@ GEXTEND Gram q_move_location: [ [ mv = move_location -> mv ] ] ; + as_name: + [ [ -> None + | "as"; id = ident_or_anti -> Some id + ] ] + ; + pose: + [ [ test_lpar_coloneq; "("; id = ident_or_anti; ":="; c = Constr.lconstr; ")" -> + Loc.tag ~loc:!@loc (Some id, c) + | c = Constr.constr; na = as_name -> Loc.tag ~loc:!@loc (na, c) + ] ] + ; + q_pose: + [ [ p = pose -> p ] ] + ; END (** Extension of constr syntax *) diff --git a/src/tac2core.ml b/src/tac2core.ml index 7bd0164b4d..bb6578090d 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -1118,6 +1118,7 @@ let () = add_expr_scope "dispatch" q_dispatch Tac2quote.of_dispatch let () = add_expr_scope "strategy" q_strategy_flag Tac2quote.of_strategy_flag let () = add_expr_scope "reference" q_reference Tac2quote.of_reference let () = add_expr_scope "move_location" q_move_location Tac2quote.of_move_location +let () = add_expr_scope "pose" q_pose Tac2quote.of_pose let () = add_expr_scope "constr_matching" q_constr_matching Tac2quote.of_constr_matching let () = add_generic_scope "constr" Pcoq.Constr.constr Tac2quote.wit_constr diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 26f96f7d72..d622aaba69 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -40,6 +40,7 @@ let q_strategy_flag = Pcoq.Gram.entry_create "tactic:q_strategy_flag" let q_constr_matching = Pcoq.Gram.entry_create "tactic:q_constr_matching" let q_hintdb = Pcoq.Gram.entry_create "tactic:q_hintdb" let q_move_location = Pcoq.Gram.entry_create "tactic:q_move_location" +let q_pose = Pcoq.Gram.entry_create "tactic:q_pose" end (** Tactic definition *) diff --git a/src/tac2entries.mli b/src/tac2entries.mli index 7c71130402..55e658884b 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -79,6 +79,7 @@ val q_strategy_flag : strategy_flag Pcoq.Gram.entry val q_constr_matching : constr_matching Pcoq.Gram.entry val q_hintdb : hintdb Pcoq.Gram.entry val q_move_location : move_location Pcoq.Gram.entry +val q_pose : pose Pcoq.Gram.entry end (** {5 Hooks} *) diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index 7d02022e07..580039afe5 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -139,3 +139,5 @@ type move_location_r = | QMoveLast type move_location = move_location_r located + +type pose = (Id.t located or_anti option * Constrexpr.constr_expr) located diff --git a/src/tac2quote.ml b/src/tac2quote.ml index f14612d58f..466c1f5094 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -370,3 +370,6 @@ let of_move_location (loc, mv) = match mv with | QMoveBefore id -> std_constructor ?loc "MoveBefore" [of_anti of_ident id] | QMoveFirst -> std_constructor ?loc "MoveFirst" [] | QMoveLast -> std_constructor ?loc "MoveLast" [] + +let of_pose p = + of_pair (fun id -> of_option (fun id -> of_anti of_ident id) id) of_open_constr p diff --git a/src/tac2quote.mli b/src/tac2quote.mli index 9f42c60042..ccb832535a 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -75,6 +75,8 @@ val of_dispatch : dispatch -> raw_tacexpr val of_strategy_flag : strategy_flag -> raw_tacexpr +val of_pose : pose -> raw_tacexpr + val of_constr_matching : constr_matching -> raw_tacexpr (** {5 Generic arguments} *) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 28bcd6a1cf..0db71fb293 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -294,13 +294,12 @@ let () = define_prim2 "tac_pose" begin fun idopt c -> Tactics.letin_tac None na c None Locusops.nowhere end -let () = define_prim4 "tac_set" begin fun ev idopt c cl -> +let () = define_prim3 "tac_set" begin fun ev p cl -> let ev = Value.to_bool ev in - let na = to_name idopt in let cl = to_clause cl in Proofview.tclEVARMAP >>= fun sigma -> - thaw c >>= fun c -> - let c = Value.to_constr c in + thaw p >>= fun p -> + let (na, c) = to_pair to_name Value.to_constr p in Tactics.letin_pat_tac ev None na (sigma, c) cl end diff --git a/tests/example2.v b/tests/example2.v index a7eb02050b..b667e19bbd 100644 --- a/tests/example2.v +++ b/tests/example2.v @@ -208,3 +208,17 @@ Proof. refine '(let x := 0 in _). eexists; exists &x; reflexivity. Qed. + +Goal True. +Proof. +pose (X := True). +constructor. +Qed. + +Goal True. +Proof. +let x := @foo in +set ($x := True) in * |-. +constructor. +Qed. + diff --git a/theories/Notations.v b/theories/Notations.v index 9b2b04f9e6..65b05113ae 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -270,6 +270,27 @@ Ltac2 Notation "apply" cl(opt(seq(keyword("in"), ident, opt(seq(keyword("as"), intropattern))))) := apply0 true false cb cl. +Ltac2 default_on_concl cl := +match cl with +| None => { Std.on_hyps := Some []; Std.on_concl := Std.AllOccurrences } +| Some cl => cl +end. + +Ltac2 pose0 ev p := + enter_h ev (fun ev ((na, p)) => Std.pose na p) p. + +Ltac2 Notation "pose" p(thunk(pose)) := + pose0 false p. + +Ltac2 Notation "epose" p(thunk(pose)) := + pose0 true p. + +Ltac2 Notation "set" p(thunk(pose)) cl(opt(clause)) := + Std.set false p (default_on_concl cl). + +Ltac2 Notation "eset" p(thunk(pose)) cl(opt(clause)) := + Std.set true p (default_on_concl cl). + Ltac2 induction0 ev ic use := let f ev use := Std.induction ev ic use in enter_h ev f use. @@ -323,12 +344,6 @@ Ltac2 Notation "inversion_clear" ids(opt(seq("in", list1(ident)))) := Std.inversion Std.FullInversionClear arg pat ids. -Ltac2 default_on_concl cl := -match cl with -| None => { Std.on_hyps := Some []; Std.on_concl := Std.AllOccurrences } -| Some cl => cl -end. - Ltac2 Notation "red" cl(opt(clause)) := Std.red (default_on_concl cl). Ltac2 Notation red := red. diff --git a/theories/Std.v b/theories/Std.v index f8b269dce6..42bd9edde0 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -135,7 +135,7 @@ Ltac2 @ external assert : constr -> (unit -> unit) option option -> intro_patter Ltac2 @ external enough : constr -> (unit -> unit) option option -> intro_pattern option -> unit := "ltac2" "tac_enough". Ltac2 @ external pose : ident option -> constr -> unit := "ltac2" "tac_pose". -Ltac2 @ external set : evar_flag -> ident option -> (unit -> constr) -> clause -> unit := "ltac2" "tac_set". +Ltac2 @ external set : evar_flag -> (unit -> ident option * constr) -> clause -> unit := "ltac2" "tac_set". Ltac2 @ external destruct : evar_flag -> induction_clause list -> constr_with_bindings option -> unit := "ltac2" "tac_induction". -- cgit v1.2.3 From 6218f06931384a38445e9d829e6782c069c3ffb4 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 14 Sep 2017 15:37:54 +0200 Subject: Introducing the remember tactic. --- src/tac2stdlib.ml | 16 ++++++++++++++++ tests/example2.v | 6 ++++++ theories/Notations.v | 20 ++++++++++++++++++++ theories/Std.v | 2 ++ 4 files changed, 44 insertions(+) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 0db71fb293..e253cc382b 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -303,6 +303,22 @@ let () = define_prim3 "tac_set" begin fun ev p cl -> Tactics.letin_pat_tac ev None na (sigma, c) cl end +let () = define_prim5 "tac_remember" begin fun ev na c eqpat cl -> + let ev = Value.to_bool ev in + let na = to_name na in + let cl = to_clause cl in + let eqpat = Value.to_option to_intro_pattern eqpat in + let eqpat = Option.default (IntroNaming IntroAnonymous) eqpat in + match eqpat with + | IntroNaming eqpat -> + Proofview.tclEVARMAP >>= fun sigma -> + thaw c >>= fun c -> + let c = Value.to_constr c in + Tactics.letin_pat_tac ev (Some (true, Loc.tag eqpat)) na (sigma, c) cl + | _ -> + Tacticals.New.tclZEROMSG (Pp.str "Invalid pattern for remember") +end + let () = define_prim3 "tac_destruct" begin fun ev ic using -> let ev = Value.to_bool ev in let ic = Value.to_list to_induction_clause ic in diff --git a/tests/example2.v b/tests/example2.v index b667e19bbd..a21f3a7f4e 100644 --- a/tests/example2.v +++ b/tests/example2.v @@ -222,3 +222,9 @@ set ($x := True) in * |-. constructor. Qed. +Goal 0 = 0. +Proof. +remember 0 as n eqn: foo at 1. +rewrite foo. +reflexivity. +Qed. diff --git a/theories/Notations.v b/theories/Notations.v index 65b05113ae..1b5792e051 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -291,6 +291,26 @@ Ltac2 Notation "set" p(thunk(pose)) cl(opt(clause)) := Ltac2 Notation "eset" p(thunk(pose)) cl(opt(clause)) := Std.set true p (default_on_concl cl). +Ltac2 default_everywhere cl := +match cl with +| None => { Std.on_hyps := None; Std.on_concl := Std.AllOccurrences } +| Some cl => cl +end. + +Ltac2 Notation "remember" + c(thunk(open_constr)) + na(opt(seq("as", ident))) + pat(opt(seq("eqn", ":", intropattern))) + cl(opt(clause)) := + Std.remember false na c pat (default_everywhere cl). + +Ltac2 Notation "eremember" + c(thunk(open_constr)) + na(opt(seq("as", ident))) + pat(opt(seq("eqn", ":", intropattern))) + cl(opt(clause)) := + Std.remember true na c pat (default_everywhere cl). + Ltac2 induction0 ev ic use := let f ev use := Std.induction ev ic use in enter_h ev f use. diff --git a/theories/Std.v b/theories/Std.v index 42bd9edde0..e938bc24b1 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -137,6 +137,8 @@ Ltac2 @ external enough : constr -> (unit -> unit) option option -> intro_patter Ltac2 @ external pose : ident option -> constr -> unit := "ltac2" "tac_pose". Ltac2 @ external set : evar_flag -> (unit -> ident option * constr) -> clause -> unit := "ltac2" "tac_set". +Ltac2 @ external remember : evar_flag -> ident option -> (unit -> constr) -> intro_pattern option -> clause -> unit := "ltac2" "tac_remember". + Ltac2 @ external destruct : evar_flag -> induction_clause list -> constr_with_bindings option -> unit := "ltac2" "tac_induction". -- cgit v1.2.3 From 7cee394fc0c6a7a28def2222be0289d6083f47c2 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 14 Sep 2017 16:07:47 +0200 Subject: Explicit arity for closures. --- src/tac2core.ml | 33 +++++++++++++-------------- src/tac2expr.mli | 6 ++++- src/tac2ffi.ml | 33 +++++++++++++++++++++++++++ src/tac2ffi.mli | 10 +++++++++ src/tac2interp.ml | 31 +++++++++++-------------- src/tac2stdlib.ml | 67 +++++++++++++++++-------------------------------------- 6 files changed, 97 insertions(+), 83 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index bb6578090d..ea65066d74 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -101,7 +101,7 @@ let err_matchfailure = (** Helper functions *) -let thaw f = f [v_unit] +let thaw f = Tac2ffi.apply f [v_unit] let fatal_flag : unit Exninfo.t = Exninfo.make () let fatal_info = Exninfo.add Exninfo.null fatal_flag () @@ -150,24 +150,21 @@ let pf_apply f = (** Primitives *) -let define0 name f = Tac2env.define_primitive (pname name) begin fun arg -> match arg with -| [_] -> f -| _ -> assert false -end +let define_primitive name arity f = + Tac2env.define_primitive (pname name) (MLTactic (arity, f)) + +let define0 name f = define_primitive name OneAty (fun _ -> f) -let define1 name r0 f = Tac2env.define_primitive (pname name) begin fun arg -> match arg with -| [x] -> f (r0.Value.r_to x) -| _ -> assert false +let define1 name r0 f = define_primitive name OneAty begin fun x -> + f (r0.Value.r_to x) end -let define2 name r0 r1 f = Tac2env.define_primitive (pname name) begin fun arg -> match arg with -| [x; y] -> f (r0.Value.r_to x) (r1.Value.r_to y) -| _ -> assert false +let define2 name r0 r1 f = define_primitive name (AddAty OneAty) begin fun x y -> + f (r0.Value.r_to x) (r1.Value.r_to y) end -let define3 name r0 r1 r2 f = Tac2env.define_primitive (pname name) begin fun arg -> match arg with -| [x; y; z] -> f (r0.Value.r_to x) (r1.Value.r_to y) (r2.Value.r_to z) -| _ -> assert false +let define3 name r0 r1 r2 f = define_primitive name (AddAty (AddAty OneAty)) begin fun x y z -> + f (r0.Value.r_to x) (r1.Value.r_to y) (r2.Value.r_to z) end (** Printing *) @@ -588,7 +585,7 @@ end (** (unit -> 'a) -> (exn -> 'a) -> 'a *) let () = define2 "plus" closure closure begin fun x k -> - Proofview.tclOR (thaw x) (fun e -> k [Value.of_exn e]) + Proofview.tclOR (thaw x) (fun e -> Tac2ffi.apply k [Value.of_exn e]) end (** (unit -> 'a) -> 'a *) @@ -620,13 +617,13 @@ end let () = define1 "case" closure begin fun f -> Proofview.tclCASE (thaw f) >>= begin function | Proofview.Next (x, k) -> - let k = function + let k = Tac2ffi.abstract 1 begin function | [e] -> let (e, info) = Value.to_exn e in set_bt info >>= fun info -> k (e, info) | _ -> assert false - in + end in return (ValBlk (0, [| Value.of_tuple [| x; Value.of_closure k |] |])) | Proofview.Fail e -> return (ValBlk (1, [| Value.of_exn e |])) end @@ -705,7 +702,7 @@ let () = define2 "with_holes" closure closure begin fun x f -> thaw x >>= fun ans -> Proofview.tclEVARMAP >>= fun sigma -> Proofview.Unsafe.tclEVARS sigma0 >>= fun () -> - Tacticals.New.tclWITHHOLES false (f [ans]) sigma + Tacticals.New.tclWITHHOLES false (Tac2ffi.apply f [ans]) sigma end let () = define1 "progress" closure begin fun f -> diff --git a/src/tac2expr.mli b/src/tac2expr.mli index 0b9923f7b9..bbe127e94d 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -188,6 +188,10 @@ type frame = type backtrace = frame list +type ('a, _) arity = +| OneAty : ('a, 'a -> 'a Proofview.tactic) arity +| AddAty : ('a, 'b) arity -> ('a, 'a -> 'b) arity + type valexpr = | ValInt of int (** Immediate integers *) @@ -202,7 +206,7 @@ type valexpr = | ValExt : 'a Tac2dyn.Val.tag * 'a -> valexpr (** Arbitrary data *) -and ml_tactic = valexpr list -> valexpr Proofview.tactic +and ml_tactic = MLTactic : (valexpr, 'v) arity * 'v -> ml_tactic type environment = { env_ist : valexpr Id.Map.t; diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index 16c17cce62..b4191acd60 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -12,6 +12,7 @@ open Globnames open Genarg open Tac2dyn open Tac2expr +open Proofview.Notations type 'a repr = { r_of : 'a -> valexpr; @@ -248,3 +249,35 @@ let reference = { r_to = to_reference; r_id = false; } + +let rec apply : type a. (valexpr, a) arity -> a -> valexpr list -> valexpr Proofview.tactic = + fun arity f args -> match args, arity with + | [], arity -> Proofview.tclUNIT (ValCls (MLTactic (arity, f))) + (** A few hardcoded cases for efficiency *) + | [a0], OneAty -> f a0 + | [a0; a1], AddAty OneAty -> f a0 a1 + | [a0; a1; a2], AddAty (AddAty OneAty) -> f a0 a1 a2 + | [a0; a1; a2; a3], AddAty (AddAty (AddAty OneAty)) -> f a0 a1 a2 a3 + (** Generic cases *) + | a :: args, OneAty -> + f a >>= fun f -> + let MLTactic (arity, f) = to_closure f in + apply arity f args + | a :: args, AddAty arity -> + apply arity (f a) args + +let apply (MLTactic (arity, f)) args = apply arity f args + +type n_closure = +| NClosure : (valexpr, 'a) arity * (valexpr list -> 'a) -> n_closure + +let rec abstract n f = + if Int.equal n 1 then NClosure (OneAty, fun accu v -> f (List.rev (v :: accu))) + else + let NClosure (arity, fe) = abstract (n - 1) f in + NClosure (AddAty arity, fun accu v -> fe (v :: accu)) + +let abstract n f = + let () = assert (n > 0) in + let NClosure (arity, f) = abstract n f in + MLTactic (arity, f []) diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index 05f4d210ab..489531671c 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -119,6 +119,16 @@ val val_exn : Exninfo.iexn Tac2dyn.Val.tag (** Toplevel representation of OCaml exceptions. Invariant: no [LtacError] should be put into a value with tag [val_exn]. *) +(** Closures *) + +val apply : ml_tactic -> valexpr list -> valexpr Proofview.tactic +(** Given a closure, apply it to some arguments. Handling of argument mismatches + is done automatically, i.e. in case of over or under-application. *) + +val abstract : int -> (valexpr list -> valexpr Proofview.tactic) -> ml_tactic +(** Turn a fixed-arity function into a closure. The inner function is guaranteed + to be applied to a list whose size is the integer argument. *) + (** Exception *) exception LtacError of KerName.t * valexpr array diff --git a/src/tac2interp.ml b/src/tac2interp.ml index 19efeb7669..815fdffe0f 100644 --- a/src/tac2interp.ml +++ b/src/tac2interp.ml @@ -88,7 +88,7 @@ let rec interp (ist : environment) = function | GTacApp (f, args) -> interp ist f >>= fun f -> Proofview.Monad.List.map (fun e -> interp ist e) args >>= fun args -> - Tac2ffi.to_closure f args + Tac2ffi.apply (Tac2ffi.to_closure f) args | GTacLet (false, el, e) -> let fold accu (na, e) = interp ist e >>= fun e -> @@ -133,28 +133,23 @@ let rec interp (ist : environment) = function return (ValOpn (kn, Array.of_list el)) | GTacPrm (ml, el) -> Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> - with_frame (FrPrim ml) (Tac2env.interp_primitive ml el) + with_frame (FrPrim ml) (Tac2ffi.apply (Tac2env.interp_primitive ml) el) | GTacExt (tag, e) -> let tpe = Tac2env.interp_ml_object tag in with_frame (FrExtn (tag, e)) (tpe.Tac2env.ml_interp ist e) -and interp_app f = (); fun args -> match f with -| { clos_env = ist; clos_var = ids; clos_exp = e; clos_ref = kn } -> - let frame = match kn with - | None -> FrAnon e - | Some kn -> FrLtac kn +and interp_app f = + let ans = fun args -> + let { clos_env = ist; clos_var = ids; clos_exp = e; clos_ref = kn } = f in + let frame = match kn with + | None -> FrAnon e + | Some kn -> FrLtac kn + in + let ist = { env_ist = ist } in + let ist = List.fold_left2 push_name ist ids args in + with_frame frame (interp ist e) in - let rec push ist ids args = match ids, args with - | [], [] -> with_frame frame (interp ist e) - | [], _ :: _ -> - with_frame frame (interp ist e) >>= fun f -> Tac2ffi.to_closure f args - | _ :: _, [] -> - let cls = { clos_ref = kn; clos_env = ist.env_ist; clos_var = ids; clos_exp = e } in - let f = interp_app cls in - return (ValCls f) - | id :: ids, arg :: args -> push (push_name ist id arg) ids args - in - push { env_ist = ist } ids args + Tac2ffi.abstract (List.length f.clos_var) ans and interp_case ist e cse0 cse1 = match e with | ValInt n -> interp ist cse0.(n) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index e253cc382b..6512510f0a 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -18,7 +18,7 @@ module Value = Tac2ffi let return x = Proofview.tclUNIT x let v_unit = Value.of_unit () -let thaw f = Value.to_closure f [v_unit] +let thaw f = Tac2ffi.apply (Value.to_closure f) [v_unit] let to_pair f g = function | ValBlk (0, [| x; y |]) -> (f x, g y) @@ -189,46 +189,28 @@ let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } let lift tac = tac <*> return v_unit let define_prim0 name tac = - let tac arg = match arg with - | [_] -> lift tac - | _ -> assert false - in - Tac2env.define_primitive (pname name) tac + let tac _ = lift tac in + Tac2env.define_primitive (pname name) (MLTactic (OneAty, tac)) let define_prim1 name tac = - let tac arg = match arg with - | [x] -> lift (tac x) - | _ -> assert false - in - Tac2env.define_primitive (pname name) tac + let tac x = lift (tac x) in + Tac2env.define_primitive (pname name) (MLTactic (OneAty, tac)) let define_prim2 name tac = - let tac arg = match arg with - | [x; y] -> lift (tac x y) - | _ -> assert false - in - Tac2env.define_primitive (pname name) tac + let tac x y = lift (tac x y) in + Tac2env.define_primitive (pname name) (MLTactic (AddAty OneAty, tac)) let define_prim3 name tac = - let tac arg = match arg with - | [x; y; z] -> lift (tac x y z) - | _ -> assert false - in - Tac2env.define_primitive (pname name) tac + let tac x y z = lift (tac x y z) in + Tac2env.define_primitive (pname name) (MLTactic (AddAty (AddAty OneAty), tac)) let define_prim4 name tac = - let tac arg = match arg with - | [x; y; z; u] -> lift (tac x y z u) - | _ -> assert false - in - Tac2env.define_primitive (pname name) tac + let tac x y z u = lift (tac x y z u) in + Tac2env.define_primitive (pname name) (MLTactic (AddAty (AddAty (AddAty OneAty)), tac)) let define_prim5 name tac = - let tac arg = match arg with - | [x; y; z; u; v] -> lift (tac x y z u v) - | _ -> assert false - in - Tac2env.define_primitive (pname name) tac + let tac x y z u v = lift (tac x y z u v) in + Tac2env.define_primitive (pname name) (MLTactic (AddAty (AddAty (AddAty (AddAty OneAty))), tac)) (** Tactics from Tacexpr *) @@ -401,26 +383,19 @@ end (** Reduction functions *) +let lift tac = tac >>= fun c -> Proofview.tclUNIT (Value.of_constr c) + let define_red1 name tac = - let tac arg = match arg with - | [x] -> tac x >>= fun c -> Proofview.tclUNIT (Value.of_constr c) - | _ -> assert false - in - Tac2env.define_primitive (pname name) tac + let tac x = lift (tac x) in + Tac2env.define_primitive (pname name) (MLTactic (OneAty, tac)) let define_red2 name tac = - let tac arg = match arg with - | [x; y] -> tac x y >>= fun c -> Proofview.tclUNIT (Value.of_constr c) - | _ -> assert false - in - Tac2env.define_primitive (pname name) tac + let tac x y = lift (tac x y) in + Tac2env.define_primitive (pname name) (MLTactic (AddAty OneAty, tac)) let define_red3 name tac = - let tac arg = match arg with - | [x; y; z] -> tac x y z >>= fun c -> Proofview.tclUNIT (Value.of_constr c) - | _ -> assert false - in - Tac2env.define_primitive (pname name) tac + let tac x y z = lift (tac x y z) in + Tac2env.define_primitive (pname name) (MLTactic (AddAty (AddAty OneAty), tac)) let () = define_red1 "eval_red" begin fun c -> let c = Value.to_constr c in -- cgit v1.2.3 From 97bcc97fab0e9c0967c7f723e24ba0f238bd94ff Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 14 Sep 2017 21:28:32 +0200 Subject: Moving valexpr definition to Tac2ffi. --- src/ltac2_plugin.mlpack | 2 +- src/tac2entries.ml | 2 +- src/tac2env.ml | 5 +++++ src/tac2env.mli | 5 +++++ src/tac2expr.mli | 24 ------------------------ src/tac2ffi.ml | 25 ++++++++++++++++++++++++- src/tac2ffi.mli | 22 ++++++++++++++++++++++ src/tac2interp.ml | 5 +++++ src/tac2interp.mli | 3 +++ src/tac2print.ml | 3 ++- src/tac2print.mli | 1 + src/tac2stdlib.ml | 1 + 12 files changed, 70 insertions(+), 28 deletions(-) diff --git a/src/ltac2_plugin.mlpack b/src/ltac2_plugin.mlpack index 70f97b9d1e..a2237f4d26 100644 --- a/src/ltac2_plugin.mlpack +++ b/src/ltac2_plugin.mlpack @@ -1,6 +1,6 @@ Tac2dyn -Tac2env Tac2ffi +Tac2env Tac2print Tac2intern Tac2interp diff --git a/src/tac2entries.ml b/src/tac2entries.ml index d622aaba69..78fe7b5bd9 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -773,7 +773,7 @@ let pr_frame = function let () = register_handler begin function | Tac2interp.LtacError (kn, args) -> let t_exn = KerName.make2 Tac2env.coq_prefix (Label.make "exn") in - let v = ValOpn (kn, args) in + let v = Tac2ffi.ValOpn (kn, args) in let t = GTypRef (Other t_exn, []) in let c = Tac2print.pr_valexpr (Global.env ()) Evd.empty v t in hov 0 (str "Uncaught Ltac2 exception:" ++ spc () ++ hov 0 c) diff --git a/src/tac2env.ml b/src/tac2env.ml index 831c6a3b42..0aa2da77ae 100644 --- a/src/tac2env.ml +++ b/src/tac2env.ml @@ -12,6 +12,7 @@ open Names open Libnames open Tac2dyn open Tac2expr +open Tac2ffi type global_data = { gdata_expr : glb_tacexpr; @@ -237,6 +238,10 @@ type 'a or_glb_tacexpr = | GlbVal of 'a | GlbTacexpr of glb_tacexpr +type environment = { + env_ist : valexpr Id.Map.t; +} + type ('a, 'b, 'r) intern_fun = Genintern.glob_sign -> 'a -> 'b * 'r glb_typexpr type ('a, 'b) ml_object = { diff --git a/src/tac2env.mli b/src/tac2env.mli index 89e22f07d5..85dba90262 100644 --- a/src/tac2env.mli +++ b/src/tac2env.mli @@ -12,6 +12,7 @@ open Libnames open Nametab open Tac2dyn open Tac2expr +open Tac2ffi (** Ltac2 global environment *) @@ -111,6 +112,10 @@ type 'a or_glb_tacexpr = type ('a, 'b, 'r) intern_fun = Genintern.glob_sign -> 'a -> 'b * 'r glb_typexpr +type environment = { + env_ist : valexpr Id.Map.t; +} + type ('a, 'b) ml_object = { ml_intern : 'r. (raw_tacexpr, glb_tacexpr, 'r) intern_fun -> ('a, 'b or_glb_tacexpr, 'r) intern_fun; ml_subst : Mod_subst.substitution -> 'b -> 'b; diff --git a/src/tac2expr.mli b/src/tac2expr.mli index bbe127e94d..c787870c65 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -187,27 +187,3 @@ type frame = | FrExtn : ('a, 'b) Tac2dyn.Arg.tag * 'b -> frame type backtrace = frame list - -type ('a, _) arity = -| OneAty : ('a, 'a -> 'a Proofview.tactic) arity -| AddAty : ('a, 'b) arity -> ('a, 'a -> 'b) arity - -type valexpr = -| ValInt of int - (** Immediate integers *) -| ValBlk of tag * valexpr array - (** Structured blocks *) -| ValStr of Bytes.t - (** Strings *) -| ValCls of ml_tactic - (** Closures *) -| ValOpn of KerName.t * valexpr array - (** Open constructors *) -| ValExt : 'a Tac2dyn.Val.tag * 'a -> valexpr - (** Arbitrary data *) - -and ml_tactic = MLTactic : (valexpr, 'v) arity * 'v -> ml_tactic - -type environment = { - env_ist : valexpr Id.Map.t; -} diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index b4191acd60..3e9842b926 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -14,6 +14,26 @@ open Tac2dyn open Tac2expr open Proofview.Notations +type ('a, _) arity = +| OneAty : ('a, 'a -> 'a Proofview.tactic) arity +| AddAty : ('a, 'b) arity -> ('a, 'a -> 'b) arity + +type valexpr = +| ValInt of int + (** Immediate integers *) +| ValBlk of tag * valexpr array + (** Structured blocks *) +| ValStr of Bytes.t + (** Strings *) +| ValCls of ml_tactic + (** Closures *) +| ValOpn of KerName.t * valexpr array + (** Open constructors *) +| ValExt : 'a Tac2dyn.Val.tag * 'a -> valexpr + (** Arbitrary data *) + +and ml_tactic = MLTactic : (valexpr, 'v) arity * 'v -> ml_tactic + type 'a repr = { r_of : 'a -> valexpr; r_to : valexpr -> 'a; @@ -166,7 +186,10 @@ let pattern = repr_ext val_pattern let internal_err = let open Names in - KerName.make2 Tac2env.coq_prefix (Label.of_id (Id.of_string "Internal")) + let coq_prefix = + MPfile (DirPath.make (List.map Id.of_string ["Init"; "Ltac2"])) + in + KerName.make2 coq_prefix (Label.of_id (Id.of_string "Internal")) (** FIXME: handle backtrace in Ltac2 exceptions *) let of_exn c = match fst c with diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index 489531671c..26e309e5fd 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -11,6 +11,28 @@ open EConstr open Tac2dyn open Tac2expr +(** {5 Toplevel values} *) + +type ('a, _) arity = +| OneAty : ('a, 'a -> 'a Proofview.tactic) arity +| AddAty : ('a, 'b) arity -> ('a, 'a -> 'b) arity + +type valexpr = +| ValInt of int + (** Immediate integers *) +| ValBlk of tag * valexpr array + (** Structured blocks *) +| ValStr of Bytes.t + (** Strings *) +| ValCls of ml_tactic + (** Closures *) +| ValOpn of KerName.t * valexpr array + (** Open constructors *) +| ValExt : 'a Tac2dyn.Val.tag * 'a -> valexpr + (** Arbitrary data *) + +and ml_tactic = MLTactic : (valexpr, 'v) arity * 'v -> ml_tactic + (** {5 Ltac2 FFI} *) type 'a repr = { diff --git a/src/tac2interp.ml b/src/tac2interp.ml index 815fdffe0f..58a3a9a4ec 100644 --- a/src/tac2interp.ml +++ b/src/tac2interp.ml @@ -13,6 +13,7 @@ open Genarg open Names open Proofview.Notations open Tac2expr +open Tac2ffi exception LtacError = Tac2ffi.LtacError @@ -42,6 +43,10 @@ let with_frame frame tac = Proofview.tclUNIT ans else tac +type environment = Tac2env.environment = { + env_ist : valexpr Id.Map.t; +} + let empty_environment = { env_ist = Id.Map.empty; } diff --git a/src/tac2interp.mli b/src/tac2interp.mli index 3acca72367..211ac95196 100644 --- a/src/tac2interp.mli +++ b/src/tac2interp.mli @@ -8,6 +8,9 @@ open Names open Tac2expr +open Tac2ffi + +type environment = Tac2env.environment val empty_environment : environment diff --git a/src/tac2print.ml b/src/tac2print.ml index 7113b01610..d39051c93e 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -12,6 +12,7 @@ open Genarg open Names open Tac2expr open Tac2env +open Tac2ffi (** Utils *) @@ -106,7 +107,7 @@ type exp_level = Tac2expr.exp_level = | E0 let pr_atom = function -| AtmInt n -> int n +| AtmInt n -> Pp.int n | AtmStr s -> qstring s let pr_name = function diff --git a/src/tac2print.mli b/src/tac2print.mli index 01abd1efb1..9b9db2937d 100644 --- a/src/tac2print.mli +++ b/src/tac2print.mli @@ -7,6 +7,7 @@ (************************************************************************) open Tac2expr +open Tac2ffi (** {5 Printing types} *) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 6512510f0a..6b3b997232 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -12,6 +12,7 @@ open Globnames open Misctypes open Genredexpr open Tac2expr +open Tac2ffi open Proofview.Notations module Value = Tac2ffi -- cgit v1.2.3 From dac0b95c77dc316a2ef65bbc3901ed7c9366e982 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 14 Sep 2017 21:36:47 +0200 Subject: Abstracting away the type of arities and ML tactics. --- src/tac2core.ml | 14 ++++++-------- src/tac2env.mli | 4 ++-- src/tac2ffi.ml | 21 ++++++++++++++------- src/tac2ffi.mli | 23 +++++++++++++---------- src/tac2stdlib.ml | 18 +++++++++--------- 5 files changed, 44 insertions(+), 36 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index ea65066d74..9e65111c0d 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -151,19 +151,19 @@ let pf_apply f = (** Primitives *) let define_primitive name arity f = - Tac2env.define_primitive (pname name) (MLTactic (arity, f)) + Tac2env.define_primitive (pname name) (mk_closure arity f) -let define0 name f = define_primitive name OneAty (fun _ -> f) +let define0 name f = define_primitive name arity_one (fun _ -> f) -let define1 name r0 f = define_primitive name OneAty begin fun x -> +let define1 name r0 f = define_primitive name arity_one begin fun x -> f (r0.Value.r_to x) end -let define2 name r0 r1 f = define_primitive name (AddAty OneAty) begin fun x y -> +let define2 name r0 r1 f = define_primitive name (arity_suc arity_one) begin fun x y -> f (r0.Value.r_to x) (r1.Value.r_to y) end -let define3 name r0 r1 r2 f = define_primitive name (AddAty (AddAty OneAty)) begin fun x y z -> +let define3 name r0 r1 r2 f = define_primitive name (arity_suc (arity_suc arity_one)) begin fun x y z -> f (r0.Value.r_to x) (r1.Value.r_to y) (r2.Value.r_to z) end @@ -617,12 +617,10 @@ end let () = define1 "case" closure begin fun f -> Proofview.tclCASE (thaw f) >>= begin function | Proofview.Next (x, k) -> - let k = Tac2ffi.abstract 1 begin function - | [e] -> + let k = Tac2ffi.mk_closure arity_one begin fun e -> let (e, info) = Value.to_exn e in set_bt info >>= fun info -> k (e, info) - | _ -> assert false end in return (ValBlk (0, [| Value.of_tuple [| x; Value.of_closure k |] |])) | Proofview.Fail e -> return (ValBlk (1, [| Value.of_exn e |])) diff --git a/src/tac2env.mli b/src/tac2env.mli index 85dba90262..b82923765d 100644 --- a/src/tac2env.mli +++ b/src/tac2env.mli @@ -101,8 +101,8 @@ val shortest_qualid_of_projection : ltac_projection -> qualid (** This state is not part of the summary, contrarily to the ones above. It is intended to be used from ML plugins to register ML-side functions. *) -val define_primitive : ml_tactic_name -> ml_tactic -> unit -val interp_primitive : ml_tactic_name -> ml_tactic +val define_primitive : ml_tactic_name -> closure -> unit +val interp_primitive : ml_tactic_name -> closure (** {5 ML primitive types} *) diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index 3e9842b926..ede4836750 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -14,9 +14,9 @@ open Tac2dyn open Tac2expr open Proofview.Notations -type ('a, _) arity = -| OneAty : ('a, 'a -> 'a Proofview.tactic) arity -| AddAty : ('a, 'b) arity -> ('a, 'a -> 'b) arity +type ('a, _) arity0 = +| OneAty : ('a, 'a -> 'a Proofview.tactic) arity0 +| AddAty : ('a, 'b) arity0 -> ('a, 'a -> 'b) arity0 type valexpr = | ValInt of int @@ -25,14 +25,21 @@ type valexpr = (** Structured blocks *) | ValStr of Bytes.t (** Strings *) -| ValCls of ml_tactic +| ValCls of closure (** Closures *) | ValOpn of KerName.t * valexpr array (** Open constructors *) | ValExt : 'a Tac2dyn.Val.tag * 'a -> valexpr (** Arbitrary data *) -and ml_tactic = MLTactic : (valexpr, 'v) arity * 'v -> ml_tactic +and closure = MLTactic : (valexpr, 'v) arity0 * 'v -> closure + +let arity_one = OneAty +let arity_suc a = AddAty a + +type 'a arity = (valexpr, 'a) arity0 + +let mk_closure arity f = MLTactic (arity, f) type 'a repr = { r_of : 'a -> valexpr; @@ -273,7 +280,7 @@ let reference = { r_id = false; } -let rec apply : type a. (valexpr, a) arity -> a -> valexpr list -> valexpr Proofview.tactic = +let rec apply : type a. a arity -> a -> valexpr list -> valexpr Proofview.tactic = fun arity f args -> match args, arity with | [], arity -> Proofview.tclUNIT (ValCls (MLTactic (arity, f))) (** A few hardcoded cases for efficiency *) @@ -292,7 +299,7 @@ let rec apply : type a. (valexpr, a) arity -> a -> valexpr list -> valexpr Proof let apply (MLTactic (arity, f)) args = apply arity f args type n_closure = -| NClosure : (valexpr, 'a) arity * (valexpr list -> 'a) -> n_closure +| NClosure : 'a arity * (valexpr list -> 'a) -> n_closure let rec abstract n f = if Int.equal n 1 then NClosure (OneAty, fun accu v -> f (List.rev (v :: accu))) diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index 26e309e5fd..ed63f2d4a6 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -13,9 +13,7 @@ open Tac2expr (** {5 Toplevel values} *) -type ('a, _) arity = -| OneAty : ('a, 'a -> 'a Proofview.tactic) arity -| AddAty : ('a, 'b) arity -> ('a, 'a -> 'b) arity +type closure type valexpr = | ValInt of int @@ -24,14 +22,19 @@ type valexpr = (** Structured blocks *) | ValStr of Bytes.t (** Strings *) -| ValCls of ml_tactic +| ValCls of closure (** Closures *) | ValOpn of KerName.t * valexpr array (** Open constructors *) | ValExt : 'a Tac2dyn.Val.tag * 'a -> valexpr (** Arbitrary data *) -and ml_tactic = MLTactic : (valexpr, 'v) arity * 'v -> ml_tactic +type 'a arity + +val arity_one : (valexpr -> valexpr Proofview.tactic) arity +val arity_suc : 'a arity -> (valexpr -> 'a) arity + +val mk_closure : 'v arity -> 'v -> closure (** {5 Ltac2 FFI} *) @@ -82,9 +85,9 @@ val of_ident : Id.t -> valexpr val to_ident : valexpr -> Id.t val ident : Id.t repr -val of_closure : ml_tactic -> valexpr -val to_closure : valexpr -> ml_tactic -val closure : ml_tactic repr +val of_closure : closure -> valexpr +val to_closure : valexpr -> closure +val closure : closure repr val block : valexpr array repr @@ -143,11 +146,11 @@ val val_exn : Exninfo.iexn Tac2dyn.Val.tag (** Closures *) -val apply : ml_tactic -> valexpr list -> valexpr Proofview.tactic +val apply : closure -> valexpr list -> valexpr Proofview.tactic (** Given a closure, apply it to some arguments. Handling of argument mismatches is done automatically, i.e. in case of over or under-application. *) -val abstract : int -> (valexpr list -> valexpr Proofview.tactic) -> ml_tactic +val abstract : int -> (valexpr list -> valexpr Proofview.tactic) -> closure (** Turn a fixed-arity function into a closure. The inner function is guaranteed to be applied to a list whose size is the integer argument. *) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 6b3b997232..e94027c899 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -191,27 +191,27 @@ let lift tac = tac <*> return v_unit let define_prim0 name tac = let tac _ = lift tac in - Tac2env.define_primitive (pname name) (MLTactic (OneAty, tac)) + Tac2env.define_primitive (pname name) (mk_closure arity_one tac) let define_prim1 name tac = let tac x = lift (tac x) in - Tac2env.define_primitive (pname name) (MLTactic (OneAty, tac)) + Tac2env.define_primitive (pname name) (mk_closure arity_one tac) let define_prim2 name tac = let tac x y = lift (tac x y) in - Tac2env.define_primitive (pname name) (MLTactic (AddAty OneAty, tac)) + Tac2env.define_primitive (pname name) (mk_closure (arity_suc arity_one) tac) let define_prim3 name tac = let tac x y z = lift (tac x y z) in - Tac2env.define_primitive (pname name) (MLTactic (AddAty (AddAty OneAty), tac)) + Tac2env.define_primitive (pname name) (mk_closure (arity_suc (arity_suc arity_one)) tac) let define_prim4 name tac = let tac x y z u = lift (tac x y z u) in - Tac2env.define_primitive (pname name) (MLTactic (AddAty (AddAty (AddAty OneAty)), tac)) + Tac2env.define_primitive (pname name) (mk_closure (arity_suc (arity_suc (arity_suc arity_one))) tac) let define_prim5 name tac = let tac x y z u v = lift (tac x y z u v) in - Tac2env.define_primitive (pname name) (MLTactic (AddAty (AddAty (AddAty (AddAty OneAty))), tac)) + Tac2env.define_primitive (pname name) (mk_closure (arity_suc (arity_suc (arity_suc (arity_suc arity_one)))) tac) (** Tactics from Tacexpr *) @@ -388,15 +388,15 @@ let lift tac = tac >>= fun c -> Proofview.tclUNIT (Value.of_constr c) let define_red1 name tac = let tac x = lift (tac x) in - Tac2env.define_primitive (pname name) (MLTactic (OneAty, tac)) + Tac2env.define_primitive (pname name) (mk_closure arity_one tac) let define_red2 name tac = let tac x y = lift (tac x y) in - Tac2env.define_primitive (pname name) (MLTactic (AddAty OneAty, tac)) + Tac2env.define_primitive (pname name) (mk_closure (arity_suc arity_one) tac) let define_red3 name tac = let tac x y z = lift (tac x y z) in - Tac2env.define_primitive (pname name) (MLTactic (AddAty (AddAty OneAty), tac)) + Tac2env.define_primitive (pname name) (mk_closure (arity_suc (arity_suc arity_one)) tac) let () = define_red1 "eval_red" begin fun c -> let c = Value.to_constr c in -- cgit v1.2.3 From 45beb72954651f3ac587faacd997a5459d122426 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 15 Sep 2017 00:35:04 +0200 Subject: Phantom type for typed closures. --- src/tac2ffi.ml | 7 +++++++ src/tac2ffi.mli | 6 ++++++ 2 files changed, 13 insertions(+) diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index ede4836750..5460643bb5 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -280,6 +280,10 @@ let reference = { r_id = false; } +type ('a, 'b) fun1 = closure + +let fun1 (r0 : 'a repr) (r1 : 'b repr) : ('a, 'b) fun1 repr = closure + let rec apply : type a. a arity -> a -> valexpr list -> valexpr Proofview.tactic = fun arity f args -> match args, arity with | [], arity -> Proofview.tclUNIT (ValCls (MLTactic (arity, f))) @@ -311,3 +315,6 @@ let abstract n f = let () = assert (n > 0) in let NClosure (arity, f) = abstract n f in MLTactic (arity, f []) + +let app_fun1 cls r0 r1 x = + apply cls [r0.r_of x] >>= fun v -> Proofview.tclUNIT (r1.r_to v) diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index ed63f2d4a6..af854e2d07 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -122,6 +122,12 @@ val of_ext : 'a Val.tag -> 'a -> valexpr val to_ext : 'a Val.tag -> valexpr -> 'a val repr_ext : 'a Val.tag -> 'a repr +type ('a, 'b) fun1 + +val app_fun1 : ('a, 'b) fun1 -> 'a repr -> 'b repr -> 'a -> 'b Proofview.tactic + +val fun1 : 'a repr -> 'b repr -> ('a, 'b) fun1 repr + val valexpr : valexpr repr (** {5 Dynamic tags} *) -- cgit v1.2.3 From 67851196e552948da9960fe32e9e9f628b349ee1 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 15 Sep 2017 00:42:13 +0200 Subject: Making Ltac2 representation of data coincide with the ML-side one. --- src/tac2stdlib.ml | 57 ++++++++--------- src/tac2tactics.ml | 178 +++++++++++++++++++++++++++++++++++++++++++++------- src/tac2tactics.mli | 79 +++++++++++++++++++---- theories/Std.v | 2 +- 4 files changed, 253 insertions(+), 63 deletions(-) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index e94027c899..07b01b1174 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -13,6 +13,7 @@ open Misctypes open Genredexpr open Tac2expr open Tac2ffi +open Tac2tactics open Proofview.Notations module Value = Tac2ffi @@ -39,7 +40,7 @@ let to_bindings = function | ValBlk (0, [| vl |]) -> ImplicitBindings (Value.to_list Value.to_constr vl) | ValBlk (1, [| vl |]) -> - ExplicitBindings ((Value.to_list (fun p -> None, to_pair to_qhyp Value.to_constr p) vl)) + ExplicitBindings ((Value.to_list (fun p -> to_pair to_qhyp Value.to_constr p) vl)) | _ -> assert false let to_constr_with_bindings = function @@ -89,8 +90,7 @@ let to_pattern_with_occs pat = to_pair Value.to_pattern (fun occ -> to_occurrences to_int_or_var occ) pat let to_constr_with_occs c = - let (c, occ) = to_pair Value.to_constr (fun occ -> to_occurrences to_int_or_var occ) c in - (occ, c) + to_pair Value.to_constr (fun occ -> to_occurrences to_int_or_var occ) c let rec to_intro_pattern = function | ValBlk (0, [| b |]) -> IntroForthcoming (Value.to_bool b) @@ -108,9 +108,11 @@ and to_intro_pattern_action = function | ValInt 0 -> IntroWildcard | ValBlk (0, [| op |]) -> IntroOrAndPattern (to_or_and_intro_pattern op) | ValBlk (1, [| inj |]) -> - let map ipat = Loc.tag (to_intro_pattern ipat) in + let map ipat = to_intro_pattern ipat in IntroInjection (Value.to_list map inj) -| ValBlk (2, [| _ |]) -> IntroApplyOn (assert false, assert false) (** TODO *) +| ValBlk (2, [| c; ipat |]) -> + let c = thaw c >>= fun c -> return (Value.to_constr c) in + IntroApplyOn (c, to_intro_pattern ipat) | ValBlk (3, [| b |]) -> IntroRewrite (Value.to_bool b) | _ -> assert false @@ -122,22 +124,21 @@ and to_or_and_intro_pattern = function | _ -> assert false and to_intro_patterns il = - let map ipat = Loc.tag (to_intro_pattern ipat) in - Value.to_list map il + Value.to_list to_intro_pattern il let to_destruction_arg = function | ValBlk (0, [| c |]) -> let c = thaw c >>= fun c -> return (to_constr_with_bindings c) in - None, ElimOnConstr c -| ValBlk (1, [| id |]) -> None, ElimOnIdent (Loc.tag (Value.to_ident id)) -| ValBlk (2, [| n |]) -> None, ElimOnAnonHyp (Value.to_int n) + ElimOnConstr c +| ValBlk (1, [| id |]) -> ElimOnIdent (Value.to_ident id) +| ValBlk (2, [| n |]) -> ElimOnAnonHyp (Value.to_int n) | _ -> assert false let to_induction_clause = function | ValBlk (0, [| arg; eqn; as_; in_ |]) -> let arg = to_destruction_arg arg in - let eqn = Value.to_option (fun p -> Loc.tag (to_intro_pattern_naming p)) eqn in - let as_ = Value.to_option (fun p -> Loc.tag (to_or_and_intro_pattern p)) as_ in + let eqn = Value.to_option to_intro_pattern_naming eqn in + let as_ = Value.to_option to_or_and_intro_pattern as_ in let in_ = Value.to_option to_clause in_ in (arg, eqn, as_, in_) | _ -> @@ -218,7 +219,7 @@ let define_prim5 name tac = let () = define_prim2 "tac_intros" begin fun ev ipat -> let ev = Value.to_bool ev in let ipat = to_intro_patterns ipat in - Tactics.intros_patterns ev ipat + Tac2tactics.intros_patterns ev ipat end let () = define_prim4 "tac_apply" begin fun adv ev cb ipat -> @@ -226,7 +227,7 @@ let () = define_prim4 "tac_apply" begin fun adv ev cb ipat -> let ev = Value.to_bool ev in let map_cb c = thaw c >>= fun c -> return (to_constr_with_bindings c) in let cb = Value.to_list map_cb cb in - let map p = Value.to_option (fun p -> Loc.tag (to_intro_pattern p)) p in + let map p = Value.to_option to_intro_pattern p in let map_ipat p = to_pair Value.to_ident map p in let ipat = Value.to_option map_ipat ipat in Tac2tactics.apply adv ev cb ipat @@ -236,13 +237,13 @@ let () = define_prim3 "tac_elim" begin fun ev c copt -> let ev = Value.to_bool ev in let c = to_constr_with_bindings c in let copt = Value.to_option to_constr_with_bindings copt in - Tactics.elim ev None c copt + Tac2tactics.elim ev c copt end let () = define_prim2 "tac_case" begin fun ev c -> let ev = Value.to_bool ev in let c = to_constr_with_bindings c in - Tactics.general_case_analysis ev None c + Tac2tactics.general_case_analysis ev c end let () = define_prim1 "tac_generalize" begin fun cl -> @@ -259,16 +260,16 @@ let () = define_prim3 "tac_assert" begin fun c tac ipat -> let c = Value.to_constr c in let of_tac t = Proofview.tclIGNORE (thaw t) in let tac = Value.to_option (fun t -> Value.to_option of_tac t) tac in - let ipat = Value.to_option (fun ipat -> Loc.tag (to_intro_pattern ipat)) ipat in - Tactics.forward true tac ipat c + let ipat = Value.to_option to_intro_pattern ipat in + Tac2tactics.forward true tac ipat c end let () = define_prim3 "tac_enough" begin fun c tac ipat -> let c = Value.to_constr c in let of_tac t = Proofview.tclIGNORE (thaw t) in let tac = Value.to_option (fun t -> Value.to_option of_tac t) tac in - let ipat = Value.to_option (fun ipat -> Loc.tag (to_intro_pattern ipat)) ipat in - Tactics.forward false tac ipat c + let ipat = Value.to_option to_intro_pattern ipat in + Tac2tactics.forward false tac ipat c end let () = define_prim2 "tac_pose" begin fun idopt c -> @@ -297,7 +298,7 @@ let () = define_prim5 "tac_remember" begin fun ev na c eqpat cl -> Proofview.tclEVARMAP >>= fun sigma -> thaw c >>= fun c -> let c = Value.to_constr c in - Tactics.letin_pat_tac ev (Some (true, Loc.tag eqpat)) na (sigma, c) cl + Tac2tactics.letin_pat_tac ev (Some (true, eqpat)) na (sigma, c) cl | _ -> Tacticals.New.tclZEROMSG (Pp.str "Invalid pattern for remember") end @@ -367,7 +368,7 @@ end let () = define_prim2 "tac_pattern" begin fun where cl -> let where = Value.to_list to_constr_with_occs where in let cl = to_clause cl in - Tactics.reduce (Pattern where) cl + Tac2tactics.pattern where cl end let () = define_prim2 "tac_vm" begin fun where cl -> @@ -476,7 +477,7 @@ end let () = define_prim4 "tac_inversion" begin fun knd arg pat ids -> let knd = to_inversion_kind knd in let arg = to_destruction_arg arg in - let pat = Value.to_option (fun ipat -> Loc.tag (to_intro_pattern ipat)) pat in + let pat = Value.to_option to_intro_pattern pat in let ids = Value.to_option (fun l -> Value.to_list Value.to_ident l) ids in Tac2tactics.inversion knd arg pat ids end @@ -523,12 +524,12 @@ end let () = define_prim2 "tac_left" begin fun ev bnd -> let ev = Value.to_bool ev in let bnd = to_bindings bnd in - Tactics.left_with_bindings ev bnd + Tac2tactics.left_with_bindings ev bnd end let () = define_prim2 "tac_right" begin fun ev bnd -> let ev = Value.to_bool ev in let bnd = to_bindings bnd in - Tactics.right_with_bindings ev bnd + Tac2tactics.right_with_bindings ev bnd end let () = define_prim1 "tac_introsuntil" begin fun h -> @@ -556,7 +557,7 @@ let () = define_prim3 "tac_constructorn" begin fun ev n bnd -> let ev = Value.to_bool ev in let n = Value.to_int n in let bnd = to_bindings bnd in - Tactics.constructor_tac ev None n bnd + Tac2tactics.constructor_tac ev None n bnd end let () = define_prim1 "tac_symmetry" begin fun cl -> @@ -567,7 +568,7 @@ end let () = define_prim2 "tac_split" begin fun ev bnd -> let ev = Value.to_bool ev in let bnd = to_bindings bnd in - Tactics.split_with_bindings ev [bnd] + Tac2tactics.split_with_bindings ev bnd end let () = define_prim1 "tac_rename" begin fun ids -> @@ -633,7 +634,7 @@ end let () = define_prim1 "tac_contradiction" begin fun c -> let c = Value.to_option to_constr_with_bindings c in - Contradiction.contradiction c + Tac2tactics.contradiction c end let () = define_prim4 "tac_autorewrite" begin fun all by ids cl -> diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index aa2ee4711a..083ec9e9d4 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -16,7 +16,41 @@ open Genredexpr open Proofview open Proofview.Notations -type destruction_arg = EConstr.constr with_bindings tactic Misctypes.destruction_arg +let return = Proofview.tclUNIT + +type explicit_bindings = (quantified_hypothesis * EConstr.t) list + +type bindings = +| ImplicitBindings of EConstr.t list +| ExplicitBindings of explicit_bindings +| NoBindings + +type constr_with_bindings = EConstr.constr * bindings + +type intro_pattern = +| IntroForthcoming of bool +| IntroNaming of intro_pattern_naming +| IntroAction of intro_pattern_action +and intro_pattern_naming = +| IntroIdentifier of Id.t +| IntroFresh of Id.t +| IntroAnonymous +and intro_pattern_action = +| IntroWildcard +| IntroOrAndPattern of or_and_intro_pattern +| IntroInjection of intro_pattern list +| IntroApplyOn of EConstr.t tactic * intro_pattern +| IntroRewrite of bool +and or_and_intro_pattern = +| IntroOrPattern of intro_pattern list list +| IntroAndPattern of intro_pattern list + +type core_destruction_arg = +| ElimOnConstr of constr_with_bindings tactic +| ElimOnIdent of Id.t +| ElimOnAnonHyp of int + +type destruction_arg = core_destruction_arg let tactic_infer_flags with_evar = { Pretyping.use_typeclasses = true; @@ -31,12 +65,57 @@ let delayed_of_tactic tac env sigma = let c, pv, _, _ = Proofview.apply env tac pv in (sigma, c) +let mk_bindings = function +| ImplicitBindings l -> Misctypes.ImplicitBindings l +| ExplicitBindings l -> + let l = List.map Loc.tag l in + Misctypes.ExplicitBindings l +| NoBindings -> Misctypes.NoBindings + +let mk_with_bindings (x, b) = (x, mk_bindings b) + +let rec mk_intro_pattern = function +| IntroForthcoming b -> Loc.tag @@ Misctypes.IntroForthcoming b +| IntroNaming ipat -> Loc.tag @@ Misctypes.IntroNaming (mk_intro_pattern_naming ipat) +| IntroAction ipat -> Loc.tag @@ Misctypes.IntroAction (mk_intro_pattern_action ipat) + +and mk_intro_pattern_naming = function +| IntroIdentifier id -> Misctypes.IntroIdentifier id +| IntroFresh id -> Misctypes.IntroFresh id +| IntroAnonymous -> Misctypes.IntroAnonymous + +and mk_intro_pattern_action = function +| IntroWildcard -> Misctypes.IntroWildcard +| IntroOrAndPattern ipat -> Misctypes.IntroOrAndPattern (mk_or_and_intro_pattern ipat) +| IntroInjection ipats -> Misctypes.IntroInjection (List.map mk_intro_pattern ipats) +| IntroApplyOn (c, ipat) -> + let c = Loc.tag @@ delayed_of_tactic c in + Misctypes.IntroApplyOn (c, mk_intro_pattern ipat) +| IntroRewrite b -> Misctypes.IntroRewrite b + +and mk_or_and_intro_pattern = function +| IntroOrPattern ipatss -> + Misctypes.IntroOrPattern (List.map (fun ipat -> List.map mk_intro_pattern ipat) ipatss) +| IntroAndPattern ipats -> + Misctypes.IntroAndPattern (List.map mk_intro_pattern ipats) + +let mk_intro_patterns ipat = List.map mk_intro_pattern ipat + +let intros_patterns ev ipat = + let ipat = mk_intro_patterns ipat in + Tactics.intro_patterns ev ipat + let apply adv ev cb cl = - let map c = None, Loc.tag (delayed_of_tactic c) in + let map c = + let c = c >>= fun c -> return (mk_with_bindings c) in + None, Loc.tag (delayed_of_tactic c) + in let cb = List.map map cb in match cl with | None -> Tactics.apply_with_delayed_bindings_gen adv ev cb - | Some (id, cl) -> Tactics.apply_delayed_in adv ev id cb cl + | Some (id, cl) -> + let cl = Option.map mk_intro_pattern cl in + Tactics.apply_delayed_in adv ev id cb cl type induction_clause = destruction_arg * @@ -44,31 +123,70 @@ type induction_clause = or_and_intro_pattern option * Locus.clause option -let map_destruction_arg = function -| ElimOnConstr c -> ElimOnConstr (delayed_of_tactic c) -| ElimOnIdent id -> ElimOnIdent id -| ElimOnAnonHyp n -> ElimOnAnonHyp n +let mk_destruction_arg = function +| ElimOnConstr c -> + let c = c >>= fun c -> return (mk_with_bindings c) in + Misctypes.ElimOnConstr (delayed_of_tactic c) +| ElimOnIdent id -> Misctypes.ElimOnIdent (Loc.tag id) +| ElimOnAnonHyp n -> Misctypes.ElimOnAnonHyp n + +let mk_induction_clause (arg, eqn, as_, occ) = + let eqn = Option.map (fun ipat -> Loc.tag @@ mk_intro_pattern_naming ipat) eqn in + let as_ = Option.map (fun ipat -> Loc.tag @@ mk_or_and_intro_pattern ipat) as_ in + ((None, mk_destruction_arg arg), (eqn, as_), occ) + +let induction_destruct isrec ev (ic : induction_clause list) using = + let ic = List.map mk_induction_clause ic in + let using = Option.map mk_with_bindings using in + Tactics.induction_destruct isrec ev (ic, using) + +let elim ev c copt = + let c = mk_with_bindings c in + let copt = Option.map mk_with_bindings copt in + Tactics.elim ev None c copt -let map_induction_clause ((clear, arg), eqn, as_, occ) = - ((clear, map_destruction_arg arg), (eqn, as_), occ) +let general_case_analysis ev c = + let c = mk_with_bindings c in + Tactics.general_case_analysis ev None c -let induction_destruct isrec ev ic using = - let ic = List.map map_induction_clause ic in - Tactics.induction_destruct isrec ev (ic, using) +let constructor_tac ev n i bnd = + let bnd = mk_bindings bnd in + Tactics.constructor_tac ev n i bnd + +let left_with_bindings ev bnd = + let bnd = mk_bindings bnd in + Tactics.left_with_bindings ev bnd + +let right_with_bindings ev bnd = + let bnd = mk_bindings bnd in + Tactics.right_with_bindings ev bnd + +let split_with_bindings ev bnd = + let bnd = mk_bindings bnd in + Tactics.split_with_bindings ev [bnd] type rewriting = bool option * multi * - EConstr.constr with_bindings tactic + constr_with_bindings tactic let rewrite ev rw cl by = let map_rw (orient, repeat, c) = + let c = c >>= fun c -> return (mk_with_bindings c) in (Option.default true orient, repeat, None, delayed_of_tactic c) in let rw = List.map map_rw rw in let by = Option.map (fun tac -> Tacticals.New.tclCOMPLETE tac, Equality.Naive) by in Equality.general_multi_rewrite ev rw cl by +let forward ev tac ipat c = + let ipat = Option.map mk_intro_pattern ipat in + Tactics.forward ev tac ipat c + +let letin_pat_tac ev ipat na c cl = + let ipat = Option.map (fun (b, ipat) -> (b, Loc.tag @@ mk_intro_pattern_naming ipat)) ipat in + Tactics.letin_pat_tac ev ipat na c cl + (** Ltac interface treats differently global references than other term arguments in reduction expressions. In Ltac1, this is done at parsing time. Instead, we parse indifferently any pattern and dispatch when the tactic is @@ -114,6 +232,10 @@ let unfold occs cl = Proofview.Monad.List.map map occs >>= fun occs -> Tactics.reduce (Unfold occs) cl +let pattern where cl = + let where = List.map (fun (c, occ) -> (occ, c)) where in + Tactics.reduce (Pattern where) cl + let vm where cl = let where = Option.map map_pattern_with_occs where in Tactics.reduce (CbvVm where) cl @@ -189,12 +311,13 @@ let on_destruction_arg tac ev arg = let env = Proofview.Goal.env gl in Proofview.tclEVARMAP >>= fun sigma -> c >>= fun (c, lbind) -> + let lbind = mk_bindings lbind in Proofview.tclEVARMAP >>= fun sigma' -> let flags = tactic_infer_flags ev in let (sigma', c) = Unification.finish_evar_resolution ~flags env sigma' (sigma, c) in - Proofview.tclUNIT (Some sigma', ElimOnConstr (c, lbind)) - | ElimOnIdent id -> Proofview.tclUNIT (None, ElimOnIdent id) - | ElimOnAnonHyp n -> Proofview.tclUNIT (None, ElimOnAnonHyp n) + Proofview.tclUNIT (Some sigma', Misctypes.ElimOnConstr (c, lbind)) + | ElimOnIdent id -> Proofview.tclUNIT (None, Misctypes.ElimOnIdent (Loc.tag id)) + | ElimOnAnonHyp n -> Proofview.tclUNIT (None, Misctypes.ElimOnAnonHyp n) in arg >>= fun (sigma', arg) -> let arg = Some (clear, arg) in @@ -204,9 +327,13 @@ let on_destruction_arg tac ev arg = Tacticals.New.tclWITHHOLES ev (tac ev arg) sigma' end -let discriminate ev arg = on_destruction_arg Equality.discr_tac ev arg +let discriminate ev arg = + let arg = Option.map (fun arg -> None, arg) arg in + on_destruction_arg Equality.discr_tac ev arg let injection ev ipat arg = + let arg = Option.map (fun arg -> None, arg) arg in + let ipat = Option.map mk_intro_patterns ipat in let tac ev arg = Equality.injClause ipat ev arg in on_destruction_arg tac ev arg @@ -262,25 +389,30 @@ let inversion knd arg pat ids = in begin match pat with | None -> Proofview.tclUNIT None - | Some (_, IntroAction (IntroOrAndPattern p)) -> - Proofview.tclUNIT (Some (Loc.tag p)) + | Some (IntroAction (IntroOrAndPattern p)) -> + Proofview.tclUNIT (Some (Loc.tag @@ mk_or_and_intro_pattern p)) | Some _ -> Tacticals.New.tclZEROMSG (str "Inversion only accept disjunctive patterns") end >>= fun pat -> let inversion _ arg = begin match arg with | None -> assert false - | Some (_, ElimOnAnonHyp n) -> + | Some (_, Misctypes.ElimOnAnonHyp n) -> Inv.inv_clause knd pat ids (AnonHyp n) - | Some (_, ElimOnIdent (_, id)) -> + | Some (_, Misctypes.ElimOnIdent (_, id)) -> Inv.inv_clause knd pat ids (NamedHyp id) - | Some (_, ElimOnConstr c) -> + | Some (_, Misctypes.ElimOnConstr c) -> + let open Misctypes in let anon = Loc.tag @@ IntroNaming IntroAnonymous in Tactics.specialize c (Some anon) >>= fun () -> Tacticals.New.onLastHypId (fun id -> Inv.inv_clause knd pat ids (NamedHyp id)) end in - on_destruction_arg inversion true (Some arg) + on_destruction_arg inversion true (Some (None, arg)) + +let contradiction c = + let c = Option.map mk_with_bindings c in + Contradiction.contradiction c (** Firstorder *) diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli index f6825d84aa..5fdd1c39bc 100644 --- a/src/tac2tactics.mli +++ b/src/tac2tactics.mli @@ -16,13 +16,41 @@ open Misctypes open Tactypes open Proofview -type destruction_arg = EConstr.constr with_bindings tactic Misctypes.destruction_arg - -(** Local reimplementations of tactics variants from Coq *) - -val apply : advanced_flag -> evars_flag -> - EConstr.constr with_bindings tactic list -> - (Id.t * intro_pattern option) option -> unit tactic +(** Redefinition of Ltac1 data structures because of impedance mismatch *) + +type explicit_bindings = (quantified_hypothesis * EConstr.t) list + +type bindings = +| ImplicitBindings of EConstr.t list +| ExplicitBindings of explicit_bindings +| NoBindings + +type constr_with_bindings = EConstr.constr * bindings + +type core_destruction_arg = +| ElimOnConstr of constr_with_bindings tactic +| ElimOnIdent of Id.t +| ElimOnAnonHyp of int + +type destruction_arg = core_destruction_arg + +type intro_pattern = +| IntroForthcoming of bool +| IntroNaming of intro_pattern_naming +| IntroAction of intro_pattern_action +and intro_pattern_naming = +| IntroIdentifier of Id.t +| IntroFresh of Id.t +| IntroAnonymous +and intro_pattern_action = +| IntroWildcard +| IntroOrAndPattern of or_and_intro_pattern +| IntroInjection of intro_pattern list +| IntroApplyOn of EConstr.t tactic * intro_pattern +| IntroRewrite of bool +and or_and_intro_pattern = +| IntroOrPattern of intro_pattern list list +| IntroAndPattern of intro_pattern list type induction_clause = destruction_arg * @@ -30,17 +58,42 @@ type induction_clause = or_and_intro_pattern option * clause option -val induction_destruct : rec_flag -> evars_flag -> - induction_clause list -> EConstr.constr with_bindings option -> unit tactic - type rewriting = bool option * multi * - EConstr.constr with_bindings tactic + constr_with_bindings tactic + +(** Local reimplementations of tactics variants from Coq *) + +val intros_patterns : evars_flag -> intro_pattern list -> unit tactic + +val apply : advanced_flag -> evars_flag -> + constr_with_bindings tactic list -> + (Id.t * intro_pattern option) option -> unit tactic + +val induction_destruct : rec_flag -> evars_flag -> + induction_clause list -> constr_with_bindings option -> unit tactic + +val elim : evars_flag -> constr_with_bindings -> constr_with_bindings option -> + unit tactic + +val general_case_analysis : evars_flag -> constr_with_bindings -> unit tactic + +val constructor_tac : evars_flag -> int option -> int -> bindings -> unit tactic + +val left_with_bindings : evars_flag -> bindings -> unit tactic +val right_with_bindings : evars_flag -> bindings -> unit tactic +val split_with_bindings : evars_flag -> bindings -> unit tactic val rewrite : evars_flag -> rewriting list -> clause -> unit tactic option -> unit tactic +val forward : bool -> unit tactic option option -> + intro_pattern option -> constr -> unit tactic + +val letin_pat_tac : evars_flag -> (bool * intro_pattern_naming) option -> + Name.t -> (Evd.evar_map * constr) -> clause -> unit tactic + val simpl : global_reference glob_red_flag -> (Pattern.constr_pattern * occurrences_expr) option -> clause -> unit tactic @@ -52,6 +105,8 @@ val lazy_ : global_reference glob_red_flag -> clause -> unit tactic val unfold : (global_reference * occurrences_expr) list -> clause -> unit tactic +val pattern : (constr * occurrences_expr) list -> clause -> unit tactic + val vm : (Pattern.constr_pattern * occurrences_expr) option -> clause -> unit tactic val native : (Pattern.constr_pattern * occurrences_expr) option -> clause -> unit tactic @@ -102,4 +157,6 @@ val typeclasses_eauto : Class_tactics.search_strategy option -> int option -> val inversion : inversion_kind -> destruction_arg -> intro_pattern option -> Id.t list option -> unit tactic +val contradiction : constr_with_bindings option -> unit tactic + val firstorder : unit Proofview.tactic option -> global_reference list -> Id.t list -> unit tactic diff --git a/theories/Std.v b/theories/Std.v index e938bc24b1..2eab172432 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -67,7 +67,7 @@ with intro_pattern_action := [ | IntroWildcard | IntroOrAndPattern (or_and_intro_pattern) | IntroInjection (intro_pattern list) -| IntroApplyOn ((constr * intro_pattern) not_implemented) (* Not Implemented yet *) +| IntroApplyOn ((unit -> constr), intro_pattern) | IntroRewrite (bool) ] with or_and_intro_pattern := [ -- cgit v1.2.3 From 3f734c5f5338feb491a6ca021e8b5a578f95c88b Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 26 Sep 2017 17:52:34 +0200 Subject: Slightly more straightforward notation for (e)apply. --- tests/example2.v | 7 +++++++ theories/Notations.v | 16 ++-------------- 2 files changed, 9 insertions(+), 14 deletions(-) diff --git a/tests/example2.v b/tests/example2.v index a21f3a7f4e..378abb86a8 100644 --- a/tests/example2.v +++ b/tests/example2.v @@ -44,6 +44,13 @@ apply &H with (m := 0). split. Qed. +Goal forall (P : nat -> Prop), (forall n m, n = m -> P n) -> (0 = 1) -> P 0. +Proof. +intros P H e. +apply &H with (m := 1) in e. +exact e. +Qed. + Goal forall (P : nat -> Prop), (forall n m, n = m -> P n) -> P 0. Proof. intros P H. diff --git a/theories/Notations.v b/theories/Notations.v index 1b5792e051..d4520dbdfd 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -246,28 +246,16 @@ Ltac2 Notation "eelim" c(thunk(constr)) bnd(thunk(with_bindings)) elim0 true c bnd use. Ltac2 apply0 adv ev cb cl := - let cl := match cl with - | None => None - | Some p => - let ((_, id, ipat)) := p in - let p := match ipat with - | None => None - | Some p => - let ((_, ipat)) := p in - Some ipat - end in - Some (id, p) - end in Std.apply adv ev cb cl. Ltac2 Notation "eapply" cb(list1(thunk(seq(constr, with_bindings)), ",")) - cl(opt(seq(keyword("in"), ident, opt(seq(keyword("as"), intropattern))))) := + cl(opt(seq("in", ident, opt(seq("as", intropattern))))) := apply0 true true cb cl. Ltac2 Notation "apply" cb(list1(thunk(seq(constr, with_bindings)), ",")) - cl(opt(seq(keyword("in"), ident, opt(seq(keyword("as"), intropattern))))) := + cl(opt(seq("in", ident, opt(seq("as", intropattern))))) := apply0 true false cb cl. Ltac2 default_on_concl cl := -- cgit v1.2.3 From 310ed15a1dd4d33246d8b331133fb7a8e7c1f4e3 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 26 Sep 2017 18:12:23 +0200 Subject: Small language of combinators for lookaheads in parsing. --- src/g_ltac2.ml4 | 106 +++++++++++++++++++++++--------------------------------- 1 file changed, 43 insertions(+), 63 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index dfd586d5ef..e6921e2f6c 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -19,81 +19,61 @@ open Ltac_plugin let err () = raise Stream.Failure +type lookahead = int -> Tok.t Stream.t -> int option + +let entry_of_lookahead s (lk : lookahead) = + let run strm = match lk 0 strm with None -> err () | Some _ -> () in + Gram.Entry.of_parser s run + +let (>>) (lk1 : lookahead) lk2 n strm = match lk1 n strm with +| None -> None +| Some n -> lk2 n strm + +let (<+>) (lk1 : lookahead) lk2 n strm = match lk1 n strm with +| None -> lk2 n strm +| Some n -> Some n + +let lk_kw kw n strm = match stream_nth n strm with +| KEYWORD kw' | IDENT kw' -> if String.equal kw kw' then Some (n + 1) else None +| _ -> None + +let lk_ident n strm = match stream_nth n strm with +| IDENT _ -> Some (n + 1) +| _ -> None + +let lk_int n strm = match stream_nth n strm with +| INT _ -> Some (n + 1) +| _ -> None + +let lk_ident_or_anti = lk_ident <+> (lk_kw "$" >> lk_ident) + (* lookahead for (x:=t), (?x:=t) and (1:=t) *) let test_lpar_idnum_coloneq = - Gram.Entry.of_parser "test_lpar_idnum_coloneq" - (fun strm -> - match stream_nth 0 strm with - | KEYWORD "(" -> - (match stream_nth 1 strm with - | IDENT _ | INT _ -> - (match stream_nth 2 strm with - | KEYWORD ":=" -> () - | _ -> err ()) - | KEYWORD "$" -> - (match stream_nth 2 strm with - | IDENT _ -> - (match stream_nth 3 strm with - | KEYWORD ":=" -> () - | _ -> err ()) - | _ -> err ()) - | _ -> err ()) - | _ -> err ()) + entry_of_lookahead "test_lpar_idnum_coloneq" begin + lk_kw "(" >> (lk_ident_or_anti <+> lk_int) >> lk_kw ":=" + end (* Hack to recognize "(x := t)" and "($x := t)" *) let test_lpar_coloneq = - Gram.Entry.of_parser "test_coloneq" - (fun strm -> - match stream_nth 0 strm with - | KEYWORD "(" -> - (match stream_nth 1 strm with - | IDENT _ -> - (match stream_nth 2 strm with - | KEYWORD ":=" -> () - | _ -> err ()) - | KEYWORD "$" -> - (match stream_nth 2 strm with - | IDENT _ -> - (match stream_nth 3 strm with - | KEYWORD ":=" -> () - | _ -> err ()) - | _ -> err ()) - | _ -> err ()) - | _ -> err ()) + entry_of_lookahead "test_coloneq" begin + lk_kw "(" >> lk_ident_or_anti >> lk_kw ":=" + end (* Hack to recognize "(x)" *) let test_lpar_id_rpar = - Gram.Entry.of_parser "lpar_id_coloneq" - (fun strm -> - match stream_nth 0 strm with - | KEYWORD "(" -> - (match stream_nth 1 strm with - | IDENT _ -> - (match stream_nth 2 strm with - | KEYWORD ")" -> () - | _ -> err ()) - | _ -> err ()) - | _ -> err ()) + entry_of_lookahead "test_lpar_id_coloneq" begin + lk_kw "(" >> lk_ident >> lk_kw ")" + end let test_ampersand_ident = - Gram.Entry.of_parser "test_ampersand_ident" - (fun strm -> - match stream_nth 0 strm with - | KEYWORD "&" -> - (match stream_nth 1 strm with - | IDENT _ -> () - | _ -> err ()) - | _ -> err ()) + entry_of_lookahead "test_ampersand_ident" begin + lk_kw "&" >> lk_ident + end let test_dollar_ident = - Gram.Entry.of_parser "test_dollar_ident" - (fun strm -> - match stream_nth 0 strm with - | IDENT "$" | KEYWORD "$" -> - (match stream_nth 1 strm with - | IDENT _ -> () - | _ -> err ()) - | _ -> err ()) + entry_of_lookahead "test_dollar_ident" begin + lk_kw "$" >> lk_ident + end let tac2expr = Tac2entries.Pltac.tac2expr let tac2type = Gram.entry_create "tactic:tac2type" -- cgit v1.2.3 From 940da8a791b8b1c704f28662fa2e6a8f3ddf040f Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 26 Sep 2017 18:57:18 +0200 Subject: Adding quotations for the assert family of tactics. --- src/g_ltac2.ml4 | 40 +++++++++++++++++++++++++++++++++++----- src/tac2core.ml | 1 + src/tac2entries.ml | 1 + src/tac2entries.mli | 1 + src/tac2qexpr.mli | 6 ++++++ src/tac2quote.ml | 11 +++++++++++ src/tac2quote.mli | 2 ++ src/tac2stdlib.ml | 20 ++++++++++++++------ src/tac2tactics.ml | 16 ++++++++++++++-- src/tac2tactics.mli | 6 ++++++ tests/example2.v | 25 +++++++++++++++++++++++++ theories/Notations.v | 7 +++++++ theories/Std.v | 7 ++++++- 13 files changed, 129 insertions(+), 14 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index e6921e2f6c..c92a242637 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -53,15 +53,21 @@ let test_lpar_idnum_coloneq = lk_kw "(" >> (lk_ident_or_anti <+> lk_int) >> lk_kw ":=" end +(* lookahead for (x:t), (?x:t) *) +let test_lpar_id_colon = + entry_of_lookahead "test_lpar_id_colon" begin + lk_kw "(" >> lk_ident_or_anti >> lk_kw ":" + end + (* Hack to recognize "(x := t)" and "($x := t)" *) -let test_lpar_coloneq = - entry_of_lookahead "test_coloneq" begin +let test_lpar_id_coloneq = + entry_of_lookahead "test_lpar_id_coloneq" begin lk_kw "(" >> lk_ident_or_anti >> lk_kw ":=" end (* Hack to recognize "(x)" *) let test_lpar_id_rpar = - entry_of_lookahead "test_lpar_id_coloneq" begin + entry_of_lookahead "test_lpar_id_rpar" begin lk_kw "(" >> lk_ident >> lk_kw ")" end @@ -370,7 +376,7 @@ GEXTEND Gram GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause q_rewriting q_clause q_dispatch q_occurrences q_strategy_flag q_destruction_arg q_reference q_with_bindings q_constr_matching - q_hintdb q_move_location q_pose; + q_hintdb q_move_location q_pose q_assert; anti: [ [ "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] ; @@ -709,7 +715,7 @@ GEXTEND Gram ] ] ; pose: - [ [ test_lpar_coloneq; "("; id = ident_or_anti; ":="; c = Constr.lconstr; ")" -> + [ [ test_lpar_id_coloneq; "("; id = ident_or_anti; ":="; c = Constr.lconstr; ")" -> Loc.tag ~loc:!@loc (Some id, c) | c = Constr.constr; na = as_name -> Loc.tag ~loc:!@loc (na, c) ] ] @@ -717,6 +723,30 @@ GEXTEND Gram q_pose: [ [ p = pose -> p ] ] ; + as_ipat: + [ [ "as"; ipat = simple_intropattern -> Some ipat + | -> None + ] ] + ; + by_tactic: + [ [ "by"; tac = tac2expr -> Some tac + | -> None + ] ] + ; + assertion: + [ [ test_lpar_id_coloneq; "("; id = ident_or_anti; ":="; c = Constr.lconstr; ")" -> + Loc.tag ~loc:!@loc (QAssertValue (id, c)) + | test_lpar_id_colon; "("; id = ident_or_anti; ":"; c = Constr.lconstr; ")"; tac = by_tactic -> + let loc = !@loc in + let ipat = Loc.tag ~loc @@ QIntroNaming (Loc.tag ~loc @@ QIntroIdentifier id) in + Loc.tag ~loc (QAssertType (Some ipat, c, tac)) + | c = Constr.constr; ipat = as_ipat; tac = by_tactic -> + Loc.tag ~loc:!@loc (QAssertType (ipat, c, tac)) + ] ] + ; + q_assert: + [ [ a = assertion -> a ] ] + ; END (** Extension of constr syntax *) diff --git a/src/tac2core.ml b/src/tac2core.ml index 9e65111c0d..1917fa5f66 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -1114,6 +1114,7 @@ let () = add_expr_scope "strategy" q_strategy_flag Tac2quote.of_strategy_flag let () = add_expr_scope "reference" q_reference Tac2quote.of_reference let () = add_expr_scope "move_location" q_move_location Tac2quote.of_move_location let () = add_expr_scope "pose" q_pose Tac2quote.of_pose +let () = add_expr_scope "assert" q_assert Tac2quote.of_assertion let () = add_expr_scope "constr_matching" q_constr_matching Tac2quote.of_constr_matching let () = add_generic_scope "constr" Pcoq.Constr.constr Tac2quote.wit_constr diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 78fe7b5bd9..24db19d02a 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -41,6 +41,7 @@ let q_constr_matching = Pcoq.Gram.entry_create "tactic:q_constr_matching" let q_hintdb = Pcoq.Gram.entry_create "tactic:q_hintdb" let q_move_location = Pcoq.Gram.entry_create "tactic:q_move_location" let q_pose = Pcoq.Gram.entry_create "tactic:q_pose" +let q_assert = Pcoq.Gram.entry_create "tactic:q_assert" end (** Tactic definition *) diff --git a/src/tac2entries.mli b/src/tac2entries.mli index 55e658884b..7bd512651c 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -80,6 +80,7 @@ val q_constr_matching : constr_matching Pcoq.Gram.entry val q_hintdb : hintdb Pcoq.Gram.entry val q_move_location : move_location Pcoq.Gram.entry val q_pose : pose Pcoq.Gram.entry +val q_assert : assertion Pcoq.Gram.entry end (** {5 Hooks} *) diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index 580039afe5..cb43a980de 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -141,3 +141,9 @@ type move_location_r = type move_location = move_location_r located type pose = (Id.t located or_anti option * Constrexpr.constr_expr) located + +type assertion_r = +| QAssertType of intro_pattern option * Constrexpr.constr_expr * raw_tacexpr option +| QAssertValue of Id.t located or_anti * Constrexpr.constr_expr + +type assertion = assertion_r located diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 466c1f5094..e89f37f2ba 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -373,3 +373,14 @@ let of_move_location (loc, mv) = match mv with let of_pose p = of_pair (fun id -> of_option (fun id -> of_anti of_ident id) id) of_open_constr p + +let of_assertion (loc, ast) = match ast with +| QAssertType (ipat, c, tac) -> + let ipat = of_option of_intro_pattern ipat in + let c = of_constr c in + let tac = of_option thunk tac in + std_constructor ?loc "AssertType" [ipat; c; tac] +| QAssertValue (id, c) -> + let id = of_anti of_ident id in + let c = of_constr c in + std_constructor ?loc "AssertValue" [id; c] diff --git a/src/tac2quote.mli b/src/tac2quote.mli index ccb832535a..b9cae23e63 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -77,6 +77,8 @@ val of_strategy_flag : strategy_flag -> raw_tacexpr val of_pose : pose -> raw_tacexpr +val of_assertion : assertion -> raw_tacexpr + val of_constr_matching : constr_matching -> raw_tacexpr (** {5 Generic arguments} *) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 07b01b1174..5f61081a76 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -144,6 +144,17 @@ let to_induction_clause = function | _ -> assert false +let to_assertion = function +| ValBlk (0, [| ipat; t; tac |]) -> + let to_tac t = Proofview.tclIGNORE (thaw t) in + let ipat = Value.to_option to_intro_pattern ipat in + let t = Value.to_constr t in + let tac = Value.to_option to_tac tac in + AssertType (ipat, t, tac) +| ValBlk (1, [| id; c |]) -> + AssertValue (Value.to_ident id, Value.to_constr c) +| _ -> assert false + let to_multi = function | ValBlk (0, [| n |]) -> Precisely (Value.to_int n) | ValBlk (1, [| n |]) -> UpTo (Value.to_int n) @@ -256,12 +267,9 @@ let () = define_prim1 "tac_generalize" begin fun cl -> Tactics.new_generalize_gen cl end -let () = define_prim3 "tac_assert" begin fun c tac ipat -> - let c = Value.to_constr c in - let of_tac t = Proofview.tclIGNORE (thaw t) in - let tac = Value.to_option (fun t -> Value.to_option of_tac t) tac in - let ipat = Value.to_option to_intro_pattern ipat in - Tac2tactics.forward true tac ipat c +let () = define_prim1 "tac_assert" begin fun ast -> + let ast = to_assertion ast in + Tac2tactics.assert_ ast end let () = define_prim3 "tac_enough" begin fun c tac ipat -> diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index 083ec9e9d4..a6dd4e3a9f 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -170,6 +170,10 @@ type rewriting = multi * constr_with_bindings tactic +type assertion = +| AssertType of intro_pattern option * EConstr.t * unit tactic option +| AssertValue of Id.t * EConstr.t + let rewrite ev rw cl by = let map_rw (orient, repeat, c) = let c = c >>= fun c -> return (mk_with_bindings c) in @@ -179,9 +183,17 @@ let rewrite ev rw cl by = let by = Option.map (fun tac -> Tacticals.New.tclCOMPLETE tac, Equality.Naive) by in Equality.general_multi_rewrite ev rw cl by -let forward ev tac ipat c = +let forward fst tac ipat c = + let ipat = Option.map mk_intro_pattern ipat in + Tactics.forward fst tac ipat c + +let assert_ = function +| AssertValue (id, c) -> + let ipat = Loc.tag @@ Misctypes.IntroNaming (Misctypes.IntroIdentifier id) in + Tactics.forward true None (Some ipat) c +| AssertType (ipat, c, tac) -> let ipat = Option.map mk_intro_pattern ipat in - Tactics.forward ev tac ipat c + Tactics.forward true (Some tac) ipat c let letin_pat_tac ev ipat na c cl = let ipat = Option.map (fun (b, ipat) -> (b, Loc.tag @@ mk_intro_pattern_naming ipat)) ipat in diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli index 5fdd1c39bc..8e15fb1392 100644 --- a/src/tac2tactics.mli +++ b/src/tac2tactics.mli @@ -63,6 +63,10 @@ type rewriting = multi * constr_with_bindings tactic +type assertion = +| AssertType of intro_pattern option * constr * unit tactic option +| AssertValue of Id.t * constr + (** Local reimplementations of tactics variants from Coq *) val intros_patterns : evars_flag -> intro_pattern list -> unit tactic @@ -91,6 +95,8 @@ val rewrite : val forward : bool -> unit tactic option option -> intro_pattern option -> constr -> unit tactic +val assert_ : assertion -> unit tactic + val letin_pat_tac : evars_flag -> (bool * intro_pattern_naming) option -> Name.t -> (Evd.evar_map * constr) -> clause -> unit tactic diff --git a/tests/example2.v b/tests/example2.v index 378abb86a8..20819606db 100644 --- a/tests/example2.v +++ b/tests/example2.v @@ -222,6 +222,12 @@ pose (X := True). constructor. Qed. +Goal True. +Proof. +pose True as X. +constructor. +Qed. + Goal True. Proof. let x := @foo in @@ -235,3 +241,22 @@ remember 0 as n eqn: foo at 1. rewrite foo. reflexivity. Qed. + +Goal True. +Proof. +assert (H := 0 + 0). +constructor. +Qed. + +Goal True. +Proof. +assert (exists n, n = 0) as [n Hn]. ++ exists 0; reflexivity. ++ exact I. +Qed. + +Goal True -> True. +Proof. +assert (H : 0 + 0 = 0) by reflexivity. +intros x; exact x. +Qed. diff --git a/theories/Notations.v b/theories/Notations.v index d4520dbdfd..91025ea964 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -279,6 +279,13 @@ Ltac2 Notation "set" p(thunk(pose)) cl(opt(clause)) := Ltac2 Notation "eset" p(thunk(pose)) cl(opt(clause)) := Std.set true p (default_on_concl cl). +Ltac2 assert0 ev ast := + enter_h ev (fun _ ast => Std.assert ast) ast. + +Ltac2 Notation "assert" ast(thunk(assert)) := assert0 false ast. + +Ltac2 Notation "eassert" ast(thunk(assert)) := assert0 true ast. + Ltac2 default_everywhere cl := match cl with | None => { Std.on_hyps := None; Std.on_concl := Std.AllOccurrences } diff --git a/theories/Std.v b/theories/Std.v index 2eab172432..7831baf046 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -88,6 +88,11 @@ Ltac2 Type induction_clause := { indcl_in : clause option; }. +Ltac2 Type assertion := [ +| AssertType (intro_pattern option, constr, (unit -> unit) option) +| AssertValue (ident, constr) +]. + Ltac2 Type repeat := [ | Precisely (int) | UpTo (int) @@ -131,7 +136,7 @@ Ltac2 @ external case : evar_flag -> constr_with_bindings -> unit := "ltac2" "ta Ltac2 @ external generalize : (constr * occurrences * ident option) list -> unit := "ltac2" "tac_generalize". -Ltac2 @ external assert : constr -> (unit -> unit) option option -> intro_pattern option -> unit := "ltac2" "tac_assert". +Ltac2 @ external assert : assertion -> unit := "ltac2" "tac_assert". Ltac2 @ external enough : constr -> (unit -> unit) option option -> intro_pattern option -> unit := "ltac2" "tac_enough". Ltac2 @ external pose : ident option -> constr -> unit := "ltac2" "tac_pose". -- cgit v1.2.3 From 5d208a8e1d46a57d3428ed43c195d193fc6c5b67 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 26 Sep 2017 22:47:09 +0200 Subject: Abstracting away the primitive functions on valexpr datatype. --- src/tac2core.ml | 123 +++++++++++++++++++++++++++-------------------------- src/tac2entries.ml | 2 +- src/tac2ffi.ml | 50 +++++++++++++++++++++- src/tac2ffi.mli | 19 ++++++++- src/tac2interp.ml | 66 +++++++++++++--------------- src/tac2print.ml | 18 ++++---- src/tac2stdlib.ml | 97 +++++++++++++++++++++--------------------- 7 files changed, 214 insertions(+), 161 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 1917fa5f66..9e3cefc6f5 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -53,7 +53,8 @@ end open Core -let v_unit = ValInt 0 +let v_unit = Value.of_unit () +let v_blk = Valexpr.make_block let of_name c = match c with | Anonymous -> Value.of_option Value.of_ident None @@ -82,8 +83,8 @@ let to_rec_declaration (nas, ts, cs) = Value.to_array Value.to_constr cs) let of_result f = function -| Inl c -> ValBlk (0, [|f c|]) -| Inr e -> ValBlk (1, [|Value.of_exn e|]) +| Inl c -> v_blk 0 [|f c|] +| Inr e -> v_blk 1 [|Value.of_exn e|] (** Stdlib exceptions *) @@ -209,19 +210,19 @@ end let () = define2 "array_make" int valexpr begin fun n x -> if n < 0 || n > Sys.max_array_length then throw err_outofbounds - else wrap (fun () -> ValBlk (0, Array.make n x)) + else wrap (fun () -> v_blk 0 (Array.make n x)) end -let () = define1 "array_length" block begin fun v -> - return (ValInt (Array.length v)) +let () = define1 "array_length" block begin fun (_, v) -> + return (Value.of_int (Array.length v)) end -let () = define3 "array_set" block int valexpr begin fun v n x -> +let () = define3 "array_set" block int valexpr begin fun (_, v) n x -> if n < 0 || n >= Array.length v then throw err_outofbounds else wrap_unit (fun () -> v.(n) <- x) end -let () = define2 "array_get" block int begin fun v n -> +let () = define2 "array_get" block int begin fun (_, v) n -> if n < 0 || n >= Array.length v then throw err_outofbounds else wrap (fun () -> v.(n)) end @@ -306,167 +307,167 @@ let () = define1 "constr_kind" constr begin fun c -> Proofview.tclEVARMAP >>= fun sigma -> return begin match EConstr.kind sigma c with | Rel n -> - ValBlk (0, [|Value.of_int n|]) + v_blk 0 [|Value.of_int n|] | Var id -> - ValBlk (1, [|Value.of_ident id|]) + v_blk 1 [|Value.of_ident id|] | Meta n -> - ValBlk (2, [|Value.of_int n|]) + v_blk 2 [|Value.of_int n|] | Evar (evk, args) -> - ValBlk (3, [| + v_blk 3 [| Value.of_int (Evar.repr evk); Value.of_array Value.of_constr args; - |]) + |] | Sort s -> - ValBlk (4, [|Value.of_ext Value.val_sort s|]) + v_blk 4 [|Value.of_ext Value.val_sort s|] | Cast (c, k, t) -> - ValBlk (5, [| + v_blk 5 [| Value.of_constr c; Value.of_ext Value.val_cast k; Value.of_constr t; - |]) + |] | Prod (na, t, u) -> - ValBlk (6, [| + v_blk 6 [| of_name na; Value.of_constr t; Value.of_constr u; - |]) + |] | Lambda (na, t, c) -> - ValBlk (7, [| + v_blk 7 [| of_name na; Value.of_constr t; Value.of_constr c; - |]) + |] | LetIn (na, b, t, c) -> - ValBlk (8, [| + v_blk 8 [| of_name na; Value.of_constr b; Value.of_constr t; Value.of_constr c; - |]) + |] | App (c, cl) -> - ValBlk (9, [| + v_blk 9 [| Value.of_constr c; Value.of_array Value.of_constr cl; - |]) + |] | Const (cst, u) -> - ValBlk (10, [| + v_blk 10 [| Value.of_constant cst; of_instance u; - |]) + |] | Ind (ind, u) -> - ValBlk (11, [| + v_blk 11 [| Value.of_ext Value.val_inductive ind; of_instance u; - |]) + |] | Construct (cstr, u) -> - ValBlk (12, [| + v_blk 12 [| Value.of_ext Value.val_constructor cstr; of_instance u; - |]) + |] | Case (ci, c, t, bl) -> - ValBlk (13, [| + v_blk 13 [| Value.of_ext Value.val_case ci; Value.of_constr c; Value.of_constr t; Value.of_array Value.of_constr bl; - |]) + |] | Fix ((recs, i), def) -> let (nas, ts, cs) = of_rec_declaration def in - ValBlk (14, [| + v_blk 14 [| Value.of_array Value.of_int recs; Value.of_int i; nas; ts; cs; - |]) + |] | CoFix (i, def) -> let (nas, ts, cs) = of_rec_declaration def in - ValBlk (15, [| + v_blk 15 [| Value.of_int i; nas; ts; cs; - |]) + |] | Proj (p, c) -> - ValBlk (16, [| + v_blk 16 [| Value.of_ext Value.val_projection p; Value.of_constr c; - |]) + |] end end let () = define1 "constr_make" valexpr begin fun knd -> let open Constr in - let c = match knd with - | ValBlk (0, [|n|]) -> + let c = match Tac2ffi.to_block knd with + | (0, [|n|]) -> let n = Value.to_int n in EConstr.mkRel n - | ValBlk (1, [|id|]) -> + | (1, [|id|]) -> let id = Value.to_ident id in EConstr.mkVar id - | ValBlk (2, [|n|]) -> + | (2, [|n|]) -> let n = Value.to_int n in EConstr.mkMeta n - | ValBlk (3, [|evk; args|]) -> + | (3, [|evk; args|]) -> let evk = Evar.unsafe_of_int (Value.to_int evk) in let args = Value.to_array Value.to_constr args in EConstr.mkEvar (evk, args) - | ValBlk (4, [|s|]) -> + | (4, [|s|]) -> let s = Value.to_ext Value.val_sort s in EConstr.mkSort (EConstr.Unsafe.to_sorts s) - | ValBlk (5, [|c; k; t|]) -> + | (5, [|c; k; t|]) -> let c = Value.to_constr c in let k = Value.to_ext Value.val_cast k in let t = Value.to_constr t in EConstr.mkCast (c, k, t) - | ValBlk (6, [|na; t; u|]) -> + | (6, [|na; t; u|]) -> let na = to_name na in let t = Value.to_constr t in let u = Value.to_constr u in EConstr.mkProd (na, t, u) - | ValBlk (7, [|na; t; c|]) -> + | (7, [|na; t; c|]) -> let na = to_name na in let t = Value.to_constr t in let u = Value.to_constr c in EConstr.mkLambda (na, t, u) - | ValBlk (8, [|na; b; t; c|]) -> + | (8, [|na; b; t; c|]) -> let na = to_name na in let b = Value.to_constr b in let t = Value.to_constr t in let c = Value.to_constr c in EConstr.mkLetIn (na, b, t, c) - | ValBlk (9, [|c; cl|]) -> + | (9, [|c; cl|]) -> let c = Value.to_constr c in let cl = Value.to_array Value.to_constr cl in EConstr.mkApp (c, cl) - | ValBlk (10, [|cst; u|]) -> + | (10, [|cst; u|]) -> let cst = Value.to_constant cst in let u = to_instance u in EConstr.mkConstU (cst, u) - | ValBlk (11, [|ind; u|]) -> + | (11, [|ind; u|]) -> let ind = Value.to_ext Value.val_inductive ind in let u = to_instance u in EConstr.mkIndU (ind, u) - | ValBlk (12, [|cstr; u|]) -> + | (12, [|cstr; u|]) -> let cstr = Value.to_ext Value.val_constructor cstr in let u = to_instance u in EConstr.mkConstructU (cstr, u) - | ValBlk (13, [|ci; c; t; bl|]) -> + | (13, [|ci; c; t; bl|]) -> let ci = Value.to_ext Value.val_case ci in let c = Value.to_constr c in let t = Value.to_constr t in let bl = Value.to_array Value.to_constr bl in EConstr.mkCase (ci, c, t, bl) - | ValBlk (14, [|recs; i; nas; ts; cs|]) -> + | (14, [|recs; i; nas; ts; cs|]) -> let recs = Value.to_array Value.to_int recs in let i = Value.to_int i in let def = to_rec_declaration (nas, ts, cs) in EConstr.mkFix ((recs, i), def) - | ValBlk (15, [|i; nas; ts; cs|]) -> + | (15, [|i; nas; ts; cs|]) -> let i = Value.to_int i in let def = to_rec_declaration (nas, ts, cs) in EConstr.mkCoFix (i, def) - | ValBlk (16, [|p; c|]) -> + | (16, [|p; c|]) -> let p = Value.to_ext Value.val_projection p in let c = Value.to_constr c in EConstr.mkProj (p, c) @@ -622,8 +623,8 @@ let () = define1 "case" closure begin fun f -> set_bt info >>= fun info -> k (e, info) end in - return (ValBlk (0, [| Value.of_tuple [| x; Value.of_closure k |] |])) - | Proofview.Fail e -> return (ValBlk (1, [| Value.of_exn e |])) + return (v_blk 0 [| Value.of_tuple [| x; Value.of_closure k |] |]) + | Proofview.Fail e -> return (v_blk 1 [| Value.of_exn e |]) end end @@ -791,7 +792,7 @@ let interp_constr flags ist c = pf_apply begin fun env sigma -> try let (sigma, c) = understand_ltac flags env sigma ist WithoutTypeConstraint c in - let c = ValExt (Value.val_constr, c) in + let c = Value.of_constr c in Proofview.Unsafe.tclEVARS sigma >>= fun () -> Proofview.tclUNIT c with e when catchable_exception e -> @@ -827,7 +828,7 @@ let () = define_ml_object Tac2quote.wit_open_constr obj let () = - let interp _ id = return (ValExt (Value.val_ident, id)) in + let interp _ id = return (Value.of_ident id) in let print _ id = str "ident:(" ++ Id.print id ++ str ")" in let obj = { ml_intern = (fun _ _ id -> GlbVal id, gtypref t_ident); @@ -843,7 +844,7 @@ let () = GlbVal pat, gtypref t_pattern in let print env pat = str "pattern:(" ++ Printer.pr_lconstr_pattern_env env Evd.empty pat ++ str ")" in - let interp _ c = return (ValExt (Value.val_pattern, c)) in + let interp _ c = return (Value.of_pattern c) in let obj = { ml_intern = intern; ml_interp = interp; diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 24db19d02a..cd4b701ca7 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -774,7 +774,7 @@ let pr_frame = function let () = register_handler begin function | Tac2interp.LtacError (kn, args) -> let t_exn = KerName.make2 Tac2env.coq_prefix (Label.make "exn") in - let v = Tac2ffi.ValOpn (kn, args) in + let v = Tac2ffi.of_open (kn, args) in let t = GTypRef (Other t_exn, []) in let c = Tac2print.pr_valexpr (Global.env ()) Evd.empty v t in hov 0 (str "Uncaught Ltac2 exception:" ++ spc () ++ hov 0 c) diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index 5460643bb5..7960d4d45f 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -41,6 +41,35 @@ type 'a arity = (valexpr, 'a) arity0 let mk_closure arity f = MLTactic (arity, f) +module Valexpr = +struct + +type t = valexpr + +let is_int = function +| ValInt _ -> true +| ValBlk _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ -> false + +let tag v = match v with +| ValBlk (n, _) -> n +| ValInt _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ -> + CErrors.anomaly (Pp.str "Unexpected value shape") + +let field v n = match v with +| ValBlk (_, v) -> v.(n) +| ValInt _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ -> + CErrors.anomaly (Pp.str "Unexpected value shape") + +let set_field v n w = match v with +| ValBlk (_, v) -> v.(n) <- w +| ValInt _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ -> + CErrors.anomaly (Pp.str "Unexpected value shape") + +let make_block tag v = ValBlk (tag, v) +let make_int n = ValInt n + +end + type 'a repr = { r_of : 'a -> valexpr; r_to : valexpr -> 'a; @@ -251,9 +280,26 @@ let array r = { r_id = false; } +let of_block (n, args) = ValBlk (n, args) +let to_block = function +| ValBlk (n, args) -> (n, args) +| _ -> assert false + let block = { - r_of = of_tuple; - r_to = to_tuple; + r_of = of_block; + r_to = to_block; + r_id = false; +} + +let of_open (kn, args) = ValOpn (kn, args) + +let to_open = function +| ValOpn (kn, args) -> (kn, args) +| _ -> assert false + +let open_ = { + r_of = of_open; + r_to = to_open; r_id = false; } diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index af854e2d07..36743f3346 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -36,6 +36,17 @@ val arity_suc : 'a arity -> (valexpr -> 'a) arity val mk_closure : 'v arity -> 'v -> closure +module Valexpr : +sig + type t = valexpr + val is_int : t -> bool + val tag : t -> int + val field : t -> int -> t + val set_field : t -> int -> t -> unit + val make_block : int -> t array -> t + val make_int : int -> t +end + (** {5 Ltac2 FFI} *) type 'a repr = { @@ -89,7 +100,9 @@ val of_closure : closure -> valexpr val to_closure : valexpr -> closure val closure : closure repr -val block : valexpr array repr +val of_block : (int * valexpr array) -> valexpr +val to_block : valexpr -> (int * valexpr array) +val block : (int * valexpr array) repr val of_array : ('a -> valexpr) -> 'a array -> valexpr val to_array : (valexpr -> 'a) -> valexpr -> 'a array @@ -122,6 +135,10 @@ val of_ext : 'a Val.tag -> 'a -> valexpr val to_ext : 'a Val.tag -> valexpr -> 'a val repr_ext : 'a Val.tag -> 'a repr +val of_open : KerName.t * valexpr array -> valexpr +val to_open : valexpr -> KerName.t * valexpr array +val open_ : (KerName.t * valexpr array) repr + type ('a, 'b) fun1 val app_fun1 : ('a, 'b) fun1 -> 'a repr -> 'b repr -> 'a -> 'b Proofview.tactic diff --git a/src/tac2interp.ml b/src/tac2interp.ml index 58a3a9a4ec..db30f52772 100644 --- a/src/tac2interp.ml +++ b/src/tac2interp.ml @@ -80,8 +80,8 @@ let get_ref ist kn = let return = Proofview.tclUNIT let rec interp (ist : environment) = function -| GTacAtm (AtmInt n) -> return (ValInt n) -| GTacAtm (AtmStr s) -> return (ValStr (Bytes.of_string s)) +| GTacAtm (AtmInt n) -> return (Tac2ffi.of_int n) +| GTacAtm (AtmStr s) -> return (Tac2ffi.of_string (Bytes.of_string s)) | GTacVar id -> return (get_var ist id) | GTacRef kn -> let data = get_ref ist kn in @@ -89,7 +89,7 @@ let rec interp (ist : environment) = function | GTacFun (ids, e) -> let cls = { clos_ref = None; clos_env = ist.env_ist; clos_var = ids; clos_exp = e } in let f = interp_app cls in - return (ValCls f) + return (Tac2ffi.of_closure f) | GTacApp (f, args) -> interp ist f >>= fun f -> Proofview.Monad.List.map (fun e -> interp ist e) args >>= fun args -> @@ -105,7 +105,7 @@ let rec interp (ist : environment) = function let map (na, e) = match e with | GTacFun (ids, e) -> let cls = { clos_ref = None; clos_env = ist.env_ist; clos_var = ids; clos_exp = e } in - let f = ValCls (interp_app cls) in + let f = Tac2ffi.of_closure (interp_app cls) in na, cls, f | _ -> anomaly (str "Ill-formed recursive function") in @@ -119,10 +119,10 @@ let rec interp (ist : environment) = function let iter (_, e, _) = e.clos_env <- ist.env_ist in let () = List.iter iter fixs in interp ist e -| GTacCst (_, n, []) -> return (ValInt n) +| GTacCst (_, n, []) -> return (Valexpr.make_int n) | GTacCst (_, n, el) -> Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> - return (ValBlk (n, Array.of_list el)) + return (Valexpr.make_block n (Array.of_list el)) | GTacCse (e, _, cse0, cse1) -> interp ist e >>= fun e -> interp_case ist e cse0 cse1 | GTacWth { opn_match = e; opn_branch = cse; opn_default = def } -> @@ -135,7 +135,7 @@ let rec interp (ist : environment) = function interp_set ist e p r | GTacOpn (kn, el) -> Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> - return (ValOpn (kn, Array.of_list el)) + return (Tac2ffi.of_open (kn, Array.of_list el)) | GTacPrm (ml, el) -> Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> with_frame (FrPrim ml) (Tac2ffi.apply (Tac2env.interp_primitive ml) el) @@ -156,17 +156,17 @@ and interp_app f = in Tac2ffi.abstract (List.length f.clos_var) ans -and interp_case ist e cse0 cse1 = match e with -| ValInt n -> interp ist cse0.(n) -| ValBlk (n, args) -> - let (ids, e) = cse1.(n) in - let ist = CArray.fold_left2 push_name ist ids args in - interp ist e -| ValExt _ | ValStr _ | ValCls _ | ValOpn _ -> - anomaly (str "Unexpected value shape") +and interp_case ist e cse0 cse1 = + if Valexpr.is_int e then + interp ist cse0.(Tac2ffi.to_int e) + else + let (n, args) = Tac2ffi.to_block e in + let (ids, e) = cse1.(n) in + let ist = CArray.fold_left2 push_name ist ids args in + interp ist e -and interp_with ist e cse def = match e with -| ValOpn (kn, args) -> +and interp_with ist e cse def = + let (kn, args) = Tac2ffi.to_open e in let br = try Some (KNmap.find kn cse) with Not_found -> None in begin match br with | None -> @@ -178,24 +178,16 @@ and interp_with ist e cse def = match e with let ist = CArray.fold_left2 push_name ist ids args in interp ist p end -| ValInt _ | ValBlk _ | ValExt _ | ValStr _ | ValCls _ -> - anomaly (str "Unexpected value shape") - -and interp_proj ist e p = match e with -| ValBlk (_, args) -> - return args.(p) -| ValInt _ | ValExt _ | ValStr _ | ValCls _ | ValOpn _ -> - anomaly (str "Unexpected value shape") - -and interp_set ist e p r = match e with -| ValBlk (_, args) -> - let () = args.(p) <- r in - return (ValInt 0) -| ValInt _ | ValExt _ | ValStr _ | ValCls _ | ValOpn _ -> - anomaly (str "Unexpected value shape") + +and interp_proj ist e p = + return (Valexpr.field e p) + +and interp_set ist e p r = + let () = Valexpr.set_field e p r in + return (Valexpr.make_int 0) and eval_pure kn = function -| GTacAtm (AtmInt n) -> ValInt n +| GTacAtm (AtmInt n) -> Valexpr.make_int n | GTacRef kn -> let { Tac2env.gdata_expr = e } = try Tac2env.interp_global kn @@ -205,10 +197,10 @@ and eval_pure kn = function | GTacFun (na, e) -> let cls = { clos_ref = kn; clos_env = Id.Map.empty; clos_var = na; clos_exp = e } in let f = interp_app cls in - ValCls f -| GTacCst (_, n, []) -> ValInt n -| GTacCst (_, n, el) -> ValBlk (n, Array.map_of_list eval_unnamed el) -| GTacOpn (kn, el) -> ValOpn (kn, Array.map_of_list eval_unnamed el) + Tac2ffi.of_closure f +| GTacCst (_, n, []) -> Valexpr.make_int n +| GTacCst (_, n, el) -> Valexpr.make_block n (Array.map_of_list eval_unnamed el) +| GTacOpn (kn, el) -> Tac2ffi.of_open (kn, Array.map_of_list eval_unnamed el) | GTacAtm (AtmStr _) | GTacLet _ | GTacVar _ | GTacSet _ | GTacApp _ | GTacCse _ | GTacPrj _ | GTacPrm _ | GTacExt _ | GTacWth _ -> anomaly (Pp.str "Term is not a syntactical value") diff --git a/src/tac2print.ml b/src/tac2print.ml index d39051c93e..45360a61f4 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -381,29 +381,27 @@ let rec pr_valexpr env sigma v t = match kind t with (** Shouldn't happen thanks to kind *) assert false | GTydAlg alg -> - begin match v with - | ValInt n -> pr_internal_constructor kn n true - | ValBlk (n, args) -> + if Valexpr.is_int v then + pr_internal_constructor kn (Tac2ffi.to_int v) true + else + let (n, args) = Tac2ffi.to_block v in let (id, tpe) = find_constructor n false alg.galg_constructors in let knc = change_kn_label kn id in let args = pr_constrargs env sigma params args tpe in hv 2 (pr_constructor knc ++ spc () ++ str "(" ++ args ++ str ")") - | _ -> str "" - end | GTydRec rcd -> str "{}" | GTydOpn -> - begin match v with - | ValOpn (knc, [||]) -> pr_constructor knc - | ValOpn (knc, args) -> + begin match Tac2ffi.to_open v with + | (knc, [||]) -> pr_constructor knc + | (knc, args) -> let data = Tac2env.interp_constructor knc in let args = pr_constrargs env sigma params args data.Tac2env.cdata_args in hv 2 (pr_constructor knc ++ spc () ++ str "(" ++ args ++ str ")") - | _ -> str "" end end | GTypArrow _ -> str "" | GTypRef (Tuple _, tl) -> - let blk = Array.to_list (block.r_to v) in + let blk = Array.to_list (snd (block.r_to v)) in if List.length blk == List.length tl then let prs = List.map2 (fun v t -> pr_valexpr env sigma v t) blk tl in hv 2 (str "(" ++ prlist_with_sep pr_comma (fun p -> p) prs ++ str ")") diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 5f61081a76..0e0eb116b3 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -22,17 +22,17 @@ let return x = Proofview.tclUNIT x let v_unit = Value.of_unit () let thaw f = Tac2ffi.apply (Value.to_closure f) [v_unit] -let to_pair f g = function -| ValBlk (0, [| x; y |]) -> (f x, g y) +let to_pair f g v = match Value.to_tuple v with +| [| x; y |] -> (f x, g y) | _ -> assert false let to_name c = match Value.to_option Value.to_ident c with | None -> Anonymous | Some id -> Name id -let to_qhyp = function -| ValBlk (0, [| i |]) -> AnonHyp (Value.to_int i) -| ValBlk (1, [| id |]) -> NamedHyp (Value.to_ident id) +let to_qhyp v = match Value.to_block v with +| (0, [| i |]) -> AnonHyp (Value.to_int i) +| (1, [| id |]) -> NamedHyp (Value.to_ident id) | _ -> assert false let to_bindings = function @@ -43,8 +43,8 @@ let to_bindings = function ExplicitBindings ((Value.to_list (fun p -> to_pair to_qhyp Value.to_constr p) vl)) | _ -> assert false -let to_constr_with_bindings = function -| ValBlk (0, [| c; bnd |]) -> (Value.to_constr c, to_bindings bnd) +let to_constr_with_bindings v = match Value.to_tuple v with +| [| c; bnd |] -> (Value.to_constr c, to_bindings bnd) | _ -> assert false let to_int_or_var i = ArgArg (Value.to_int i) @@ -56,16 +56,16 @@ let to_occurrences f = function | ValBlk (1, [| vl |]) -> OnlyOccurrences (Value.to_list f vl) | _ -> assert false -let to_hyp_location_flag = function -| ValInt 0 -> InHyp -| ValInt 1 -> InHypTypeOnly -| ValInt 2 -> InHypValueOnly +let to_hyp_location_flag v = match Value.to_int v with +| 0 -> InHyp +| 1 -> InHypTypeOnly +| 2 -> InHypValueOnly | _ -> assert false -let to_clause = function -| ValBlk (0, [| hyps; concl |]) -> - let cast = function - | ValBlk (0, [| hyp; occ; flag |]) -> +let to_clause v = match Value.to_tuple v with +| [| hyps; concl |] -> + let cast v = match Value.to_tuple v with + | [| hyp; occ; flag |] -> ((to_occurrences to_int_or_var occ, Value.to_ident hyp), to_hyp_location_flag flag) | _ -> assert false in @@ -73,8 +73,8 @@ let to_clause = function { onhyps = hyps; concl_occs = to_occurrences to_int_or_var concl; } | _ -> assert false -let to_red_flag = function -| ValBlk (0, [| beta; iota; fix; cofix; zeta; delta; const |]) -> +let to_red_flag v = match Value.to_tuple v with +| [| beta; iota; fix; cofix; zeta; delta; const |] -> { rBeta = Value.to_bool beta; rMatch = Value.to_bool iota; @@ -92,10 +92,10 @@ let to_pattern_with_occs pat = let to_constr_with_occs c = to_pair Value.to_constr (fun occ -> to_occurrences to_int_or_var occ) c -let rec to_intro_pattern = function -| ValBlk (0, [| b |]) -> IntroForthcoming (Value.to_bool b) -| ValBlk (1, [| pat |]) -> IntroNaming (to_intro_pattern_naming pat) -| ValBlk (2, [| act |]) -> IntroAction (to_intro_pattern_action act) +let rec to_intro_pattern v = match Value.to_block v with +| (0, [| b |]) -> IntroForthcoming (Value.to_bool b) +| (1, [| pat |]) -> IntroNaming (to_intro_pattern_naming pat) +| (2, [| act |]) -> IntroAction (to_intro_pattern_action act) | _ -> assert false and to_intro_pattern_naming = function @@ -116,26 +116,26 @@ and to_intro_pattern_action = function | ValBlk (3, [| b |]) -> IntroRewrite (Value.to_bool b) | _ -> assert false -and to_or_and_intro_pattern = function -| ValBlk (0, [| ill |]) -> +and to_or_and_intro_pattern v = match Value.to_block v with +| (0, [| ill |]) -> IntroOrPattern (Value.to_list to_intro_patterns ill) -| ValBlk (1, [| il |]) -> +| (1, [| il |]) -> IntroAndPattern (to_intro_patterns il) | _ -> assert false and to_intro_patterns il = Value.to_list to_intro_pattern il -let to_destruction_arg = function -| ValBlk (0, [| c |]) -> +let to_destruction_arg v = match Value.to_block v with +| (0, [| c |]) -> let c = thaw c >>= fun c -> return (to_constr_with_bindings c) in ElimOnConstr c -| ValBlk (1, [| id |]) -> ElimOnIdent (Value.to_ident id) -| ValBlk (2, [| n |]) -> ElimOnAnonHyp (Value.to_int n) +| (1, [| id |]) -> ElimOnIdent (Value.to_ident id) +| (2, [| n |]) -> ElimOnAnonHyp (Value.to_int n) | _ -> assert false -let to_induction_clause = function -| ValBlk (0, [| arg; eqn; as_; in_ |]) -> +let to_induction_clause v = match Value.to_tuple v with +| [| arg; eqn; as_; in_ |] -> let arg = to_destruction_arg arg in let eqn = Value.to_option to_intro_pattern_naming eqn in let as_ = Value.to_option to_or_and_intro_pattern as_ in @@ -144,14 +144,14 @@ let to_induction_clause = function | _ -> assert false -let to_assertion = function -| ValBlk (0, [| ipat; t; tac |]) -> +let to_assertion v = match Value.to_block v with +| (0, [| ipat; t; tac |]) -> let to_tac t = Proofview.tclIGNORE (thaw t) in let ipat = Value.to_option to_intro_pattern ipat in let t = Value.to_constr t in let tac = Value.to_option to_tac tac in AssertType (ipat, t, tac) -| ValBlk (1, [| id; c |]) -> +| (1, [| id; c |]) -> AssertValue (Value.to_ident id, Value.to_constr c) | _ -> assert false @@ -162,30 +162,29 @@ let to_multi = function | ValInt 1 -> RepeatPlus | _ -> assert false -let to_rewriting = function -| ValBlk (0, [| orient; repeat; c |]) -> +let to_rewriting v = match Value.to_tuple v with +| [| orient; repeat; c |] -> let orient = Value.to_option Value.to_bool orient in let repeat = to_multi repeat in - (** FIXME: lost backtrace *) let c = thaw c >>= fun c -> return (to_constr_with_bindings c) in (orient, repeat, c) | _ -> assert false -let to_debug = function -| ValInt 0 -> Hints.Off -| ValInt 1 -> Hints.Info -| ValInt 2 -> Hints.Debug +let to_debug v = match Value.to_int v with +| 0 -> Hints.Off +| 1 -> Hints.Info +| 2 -> Hints.Debug | _ -> assert false -let to_strategy = function -| ValInt 0 -> Class_tactics.Bfs -| ValInt 1 -> Class_tactics.Dfs +let to_strategy v = match Value.to_int v with +| 0 -> Class_tactics.Bfs +| 1 -> Class_tactics.Dfs | _ -> assert false -let to_inversion_kind = function -| ValInt 0 -> Misctypes.SimpleInversion -| ValInt 1 -> Misctypes.FullInversion -| ValInt 2 -> Misctypes.FullInversionClear +let to_inversion_kind v = match Value.to_int v with +| 0 -> Misctypes.SimpleInversion +| 1 -> Misctypes.FullInversion +| 2 -> Misctypes.FullInversionClear | _ -> assert false let to_move_location = function @@ -258,8 +257,8 @@ let () = define_prim2 "tac_case" begin fun ev c -> end let () = define_prim1 "tac_generalize" begin fun cl -> - let cast = function - | ValBlk (0, [| c; occs; na |]) -> + let cast v = match Value.to_tuple v with + | [| c; occs; na |] -> ((to_occurrences Value.to_int occs, Value.to_constr c), to_name na) | _ -> assert false in -- cgit v1.2.3 From 65be2f00dc464493edb8031544b61db6216d453c Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 1 Oct 2017 14:42:01 +0200 Subject: Moving ML types used by Ltac2 to their proper interface. --- _CoqProject | 1 + src/tac2stdlib.ml | 14 +++++----- src/tac2tactics.ml | 51 +----------------------------------- src/tac2tactics.mli | 56 ++------------------------------------- src/tac2types.mli | 75 +++++++++++++++++++++++++++++++++++++++++++++++++++++ 5 files changed, 86 insertions(+), 111 deletions(-) create mode 100644 src/tac2types.mli diff --git a/_CoqProject b/_CoqProject index fc9df4ee3f..43e9b76991 100644 --- a/_CoqProject +++ b/_CoqProject @@ -5,6 +5,7 @@ src/tac2dyn.ml src/tac2dyn.mli src/tac2expr.mli +src/tac2types.mli src/tac2env.ml src/tac2env.mli src/tac2print.ml diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 0e0eb116b3..13f150381a 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -9,10 +9,10 @@ open Names open Locus open Globnames -open Misctypes open Genredexpr open Tac2expr open Tac2ffi +open Tac2types open Tac2tactics open Proofview.Notations @@ -47,7 +47,7 @@ let to_constr_with_bindings v = match Value.to_tuple v with | [| c; bnd |] -> (Value.to_constr c, to_bindings bnd) | _ -> assert false -let to_int_or_var i = ArgArg (Value.to_int i) +let to_int_or_var i = Misctypes.ArgArg (Value.to_int i) let to_occurrences f = function | ValInt 0 -> AllOccurrences @@ -188,10 +188,10 @@ let to_inversion_kind v = match Value.to_int v with | _ -> assert false let to_move_location = function -| ValInt 0 -> MoveFirst -| ValInt 1 -> MoveLast -| ValBlk (0, [|id|]) -> MoveAfter (Value.to_ident id) -| ValBlk (1, [|id|]) -> MoveBefore (Value.to_ident id) +| ValInt 0 -> Misctypes.MoveFirst +| ValInt 1 -> Misctypes.MoveLast +| ValBlk (0, [|id|]) -> Misctypes.MoveAfter (Value.to_ident id) +| ValBlk (1, [|id|]) -> Misctypes.MoveBefore (Value.to_ident id) | _ -> assert false (** Standard tactics sharing their implementation with Ltac1 *) @@ -502,7 +502,7 @@ end let () = define_prim2 "tac_intro" begin fun id mv -> let id = Value.to_option Value.to_ident id in let mv = Value.to_option to_move_location mv in - let mv = Option.default MoveLast mv in + let mv = Option.default Misctypes.MoveLast mv in Tactics.intro_move id mv end diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index a6dd4e3a9f..42916a9578 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -11,47 +11,13 @@ open Util open Names open Globnames open Misctypes -open Tactypes +open Tac2types open Genredexpr open Proofview open Proofview.Notations let return = Proofview.tclUNIT -type explicit_bindings = (quantified_hypothesis * EConstr.t) list - -type bindings = -| ImplicitBindings of EConstr.t list -| ExplicitBindings of explicit_bindings -| NoBindings - -type constr_with_bindings = EConstr.constr * bindings - -type intro_pattern = -| IntroForthcoming of bool -| IntroNaming of intro_pattern_naming -| IntroAction of intro_pattern_action -and intro_pattern_naming = -| IntroIdentifier of Id.t -| IntroFresh of Id.t -| IntroAnonymous -and intro_pattern_action = -| IntroWildcard -| IntroOrAndPattern of or_and_intro_pattern -| IntroInjection of intro_pattern list -| IntroApplyOn of EConstr.t tactic * intro_pattern -| IntroRewrite of bool -and or_and_intro_pattern = -| IntroOrPattern of intro_pattern list list -| IntroAndPattern of intro_pattern list - -type core_destruction_arg = -| ElimOnConstr of constr_with_bindings tactic -| ElimOnIdent of Id.t -| ElimOnAnonHyp of int - -type destruction_arg = core_destruction_arg - let tactic_infer_flags with_evar = { Pretyping.use_typeclasses = true; Pretyping.solve_unification_constraints = true; @@ -117,12 +83,6 @@ let apply adv ev cb cl = let cl = Option.map mk_intro_pattern cl in Tactics.apply_delayed_in adv ev id cb cl -type induction_clause = - destruction_arg * - intro_pattern_naming option * - or_and_intro_pattern option * - Locus.clause option - let mk_destruction_arg = function | ElimOnConstr c -> let c = c >>= fun c -> return (mk_with_bindings c) in @@ -165,15 +125,6 @@ let split_with_bindings ev bnd = let bnd = mk_bindings bnd in Tactics.split_with_bindings ev [bnd] -type rewriting = - bool option * - multi * - constr_with_bindings tactic - -type assertion = -| AssertType of intro_pattern option * EConstr.t * unit tactic option -| AssertValue of Id.t * EConstr.t - let rewrite ev rw cl by = let map_rw (orient, repeat, c) = let c = c >>= fun c -> return (mk_with_bindings c) in diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli index 8e15fb1392..842b09c22f 100644 --- a/src/tac2tactics.mli +++ b/src/tac2tactics.mli @@ -12,61 +12,9 @@ open Globnames open Tac2expr open EConstr open Genredexpr -open Misctypes -open Tactypes +open Tac2types open Proofview -(** Redefinition of Ltac1 data structures because of impedance mismatch *) - -type explicit_bindings = (quantified_hypothesis * EConstr.t) list - -type bindings = -| ImplicitBindings of EConstr.t list -| ExplicitBindings of explicit_bindings -| NoBindings - -type constr_with_bindings = EConstr.constr * bindings - -type core_destruction_arg = -| ElimOnConstr of constr_with_bindings tactic -| ElimOnIdent of Id.t -| ElimOnAnonHyp of int - -type destruction_arg = core_destruction_arg - -type intro_pattern = -| IntroForthcoming of bool -| IntroNaming of intro_pattern_naming -| IntroAction of intro_pattern_action -and intro_pattern_naming = -| IntroIdentifier of Id.t -| IntroFresh of Id.t -| IntroAnonymous -and intro_pattern_action = -| IntroWildcard -| IntroOrAndPattern of or_and_intro_pattern -| IntroInjection of intro_pattern list -| IntroApplyOn of EConstr.t tactic * intro_pattern -| IntroRewrite of bool -and or_and_intro_pattern = -| IntroOrPattern of intro_pattern list list -| IntroAndPattern of intro_pattern list - -type induction_clause = - destruction_arg * - intro_pattern_naming option * - or_and_intro_pattern option * - clause option - -type rewriting = - bool option * - multi * - constr_with_bindings tactic - -type assertion = -| AssertType of intro_pattern option * constr * unit tactic option -| AssertValue of Id.t * constr - (** Local reimplementations of tactics variants from Coq *) val intros_patterns : evars_flag -> intro_pattern list -> unit tactic @@ -161,7 +109,7 @@ val eauto : Hints.debug -> int option -> int option -> constr tactic list -> val typeclasses_eauto : Class_tactics.search_strategy option -> int option -> Id.t list option -> unit Proofview.tactic -val inversion : inversion_kind -> destruction_arg -> intro_pattern option -> Id.t list option -> unit tactic +val inversion : Misctypes.inversion_kind -> destruction_arg -> intro_pattern option -> Id.t list option -> unit tactic val contradiction : constr_with_bindings option -> unit tactic diff --git a/src/tac2types.mli b/src/tac2types.mli new file mode 100644 index 0000000000..6845de8c7c --- /dev/null +++ b/src/tac2types.mli @@ -0,0 +1,75 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* (Value.to_constr c, to_bindings bnd) | _ -> assert false -let to_int_or_var i = Misctypes.ArgArg (Value.to_int i) - -let to_occurrences f = function +let to_occurrences = function | ValInt 0 -> AllOccurrences -| ValBlk (0, [| vl |]) -> AllOccurrencesBut (Value.to_list f vl) +| ValBlk (0, [| vl |]) -> AllOccurrencesBut (Value.to_list Value.to_int vl) | ValInt 1 -> NoOccurrences -| ValBlk (1, [| vl |]) -> OnlyOccurrences (Value.to_list f vl) +| ValBlk (1, [| vl |]) -> OnlyOccurrences (Value.to_list Value.to_int vl) | _ -> assert false let to_hyp_location_flag v = match Value.to_int v with @@ -66,11 +63,11 @@ let to_clause v = match Value.to_tuple v with | [| hyps; concl |] -> let cast v = match Value.to_tuple v with | [| hyp; occ; flag |] -> - ((to_occurrences to_int_or_var occ, Value.to_ident hyp), to_hyp_location_flag flag) + (Value.to_ident hyp, to_occurrences occ, to_hyp_location_flag flag) | _ -> assert false in let hyps = Value.to_option (fun h -> Value.to_list cast h) hyps in - { onhyps = hyps; concl_occs = to_occurrences to_int_or_var concl; } + { onhyps = hyps; concl_occs = to_occurrences concl; } | _ -> assert false let to_red_flag v = match Value.to_tuple v with @@ -87,10 +84,10 @@ let to_red_flag v = match Value.to_tuple v with | _ -> assert false let to_pattern_with_occs pat = - to_pair Value.to_pattern (fun occ -> to_occurrences to_int_or_var occ) pat + to_pair Value.to_pattern to_occurrences pat let to_constr_with_occs c = - to_pair Value.to_constr (fun occ -> to_occurrences to_int_or_var occ) c + to_pair Value.to_constr to_occurrences c let rec to_intro_pattern v = match Value.to_block v with | (0, [| b |]) -> IntroForthcoming (Value.to_bool b) @@ -259,11 +256,11 @@ end let () = define_prim1 "tac_generalize" begin fun cl -> let cast v = match Value.to_tuple v with | [| c; occs; na |] -> - ((to_occurrences Value.to_int occs, Value.to_constr c), to_name na) + (Value.to_constr c, to_occurrences occs, to_name na) | _ -> assert false in let cl = Value.to_list cast cl in - Tactics.new_generalize_gen cl + Tac2tactics.generalize cl end let () = define_prim1 "tac_assert" begin fun ast -> @@ -291,7 +288,7 @@ let () = define_prim3 "tac_set" begin fun ev p cl -> Proofview.tclEVARMAP >>= fun sigma -> thaw p >>= fun p -> let (na, c) = to_pair to_name Value.to_constr p in - Tactics.letin_pat_tac ev None na (sigma, c) cl + Tac2tactics.letin_pat_tac ev None na (sigma, c) cl end let () = define_prim5 "tac_remember" begin fun ev na c eqpat cl -> @@ -326,12 +323,12 @@ end let () = define_prim1 "tac_red" begin fun cl -> let cl = to_clause cl in - Tactics.reduce (Red false) cl + Tac2tactics.reduce (Red false) cl end let () = define_prim1 "tac_hnf" begin fun cl -> let cl = to_clause cl in - Tactics.reduce Hnf cl + Tac2tactics.reduce Hnf cl end let () = define_prim3 "tac_simpl" begin fun flags where cl -> @@ -360,7 +357,7 @@ let () = define_prim2 "tac_lazy" begin fun flags cl -> end let () = define_prim2 "tac_unfold" begin fun refs cl -> - let map v = to_pair Value.to_reference (fun occ -> to_occurrences to_int_or_var occ) v in + let map v = to_pair Value.to_reference to_occurrences v in let refs = Value.to_list map refs in let cl = to_clause cl in Tac2tactics.unfold refs cl @@ -369,7 +366,7 @@ end let () = define_prim2 "tac_fold" begin fun args cl -> let args = Value.to_list Value.to_constr args in let cl = to_clause cl in - Tactics.reduce (Fold args) cl + Tac2tactics.reduce (Fold args) cl end let () = define_prim2 "tac_pattern" begin fun where cl -> @@ -442,7 +439,7 @@ let () = define_red2 "eval_lazy" begin fun flags c -> end let () = define_red2 "eval_unfold" begin fun refs c -> - let map v = to_pair Value.to_reference (fun occ -> to_occurrences to_int_or_var occ) v in + let map v = to_pair Value.to_reference to_occurrences v in let refs = Value.to_list map refs in let c = Value.to_constr c in Tac2tactics.eval_unfold refs c @@ -455,7 +452,7 @@ let () = define_red2 "eval_fold" begin fun args c -> end let () = define_red2 "eval_pattern" begin fun where c -> - let where = Value.to_list (fun p -> to_pair Value.to_constr (fun occ -> to_occurrences to_int_or_var occ) p) where in + let where = Value.to_list (fun p -> to_pair Value.to_constr to_occurrences p) where in let c = Value.to_constr c in Tac2tactics.eval_pattern where c end @@ -569,7 +566,7 @@ end let () = define_prim1 "tac_symmetry" begin fun cl -> let cl = to_clause cl in - Tactics.intros_symmetry cl + Tac2tactics.symmetry cl end let () = define_prim2 "tac_split" begin fun ev bnd -> diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index 42916a9578..b55bd5c1b8 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -67,6 +67,23 @@ and mk_or_and_intro_pattern = function let mk_intro_patterns ipat = List.map mk_intro_pattern ipat +let mk_occurrences f = function +| AllOccurrences -> Locus.AllOccurrences +| AllOccurrencesBut l -> Locus.AllOccurrencesBut (List.map f l) +| NoOccurrences -> Locus.NoOccurrences +| OnlyOccurrences l -> Locus.OnlyOccurrences (List.map f l) + +let mk_occurrences_expr occ = + mk_occurrences (fun i -> Misctypes.ArgArg i) occ + +let mk_hyp_location (id, occs, h) = + ((mk_occurrences_expr occs, id), h) + +let mk_clause cl = { + Locus.onhyps = Option.map (fun l -> List.map mk_hyp_location l) cl.onhyps; + Locus.concl_occs = mk_occurrences_expr cl.concl_occs; +} + let intros_patterns ev ipat = let ipat = mk_intro_patterns ipat in Tactics.intro_patterns ev ipat @@ -93,6 +110,7 @@ let mk_destruction_arg = function let mk_induction_clause (arg, eqn, as_, occ) = let eqn = Option.map (fun ipat -> Loc.tag @@ mk_intro_pattern_naming ipat) eqn in let as_ = Option.map (fun ipat -> Loc.tag @@ mk_or_and_intro_pattern ipat) as_ in + let occ = Option.map mk_clause occ in ((None, mk_destruction_arg arg), (eqn, as_), occ) let induction_destruct isrec ev (ic : induction_clause list) using = @@ -105,6 +123,11 @@ let elim ev c copt = let copt = Option.map mk_with_bindings copt in Tactics.elim ev None c copt +let generalize pl = + let mk_occ occs = mk_occurrences (fun i -> i) occs in + let pl = List.map (fun (c, occs, na) -> (mk_occ occs, c), na) pl in + Tactics.new_generalize_gen pl + let general_case_analysis ev c = let c = mk_with_bindings c in Tactics.general_case_analysis ev None c @@ -131,9 +154,14 @@ let rewrite ev rw cl by = (Option.default true orient, repeat, None, delayed_of_tactic c) in let rw = List.map map_rw rw in + let cl = mk_clause cl in let by = Option.map (fun tac -> Tacticals.New.tclCOMPLETE tac, Equality.Naive) by in Equality.general_multi_rewrite ev rw cl by +let symmetry cl = + let cl = mk_clause cl in + Tactics.intros_symmetry cl + let forward fst tac ipat c = let ipat = Option.map mk_intro_pattern ipat in Tactics.forward fst tac ipat c @@ -148,6 +176,7 @@ let assert_ = function let letin_pat_tac ev ipat na c cl = let ipat = Option.map (fun (b, ipat) -> (b, Loc.tag @@ mk_intro_pattern_naming ipat)) ipat in + let cl = mk_clause cl in Tactics.letin_pat_tac ev ipat na c cl (** Ltac interface treats differently global references than other term @@ -155,9 +184,9 @@ let letin_pat_tac ev ipat na c cl = Instead, we parse indifferently any pattern and dispatch when the tactic is called. *) let map_pattern_with_occs (pat, occ) = match pat with -| Pattern.PRef (ConstRef cst) -> (occ, Inl (EvalConstRef cst)) -| Pattern.PRef (VarRef id) -> (occ, Inl (EvalVarRef id)) -| _ -> (occ, Inr pat) +| Pattern.PRef (ConstRef cst) -> (mk_occurrences_expr occ, Inl (EvalConstRef cst)) +| Pattern.PRef (VarRef id) -> (mk_occurrences_expr occ, Inl (EvalVarRef id)) +| _ -> (mk_occurrences_expr occ, Inr pat) let get_evaluable_reference = function | VarRef id -> Proofview.tclUNIT (EvalVarRef id) @@ -167,44 +196,57 @@ let get_evaluable_reference = function Nametab.pr_global_env Id.Set.empty r ++ spc () ++ str "to an evaluable reference.") +let reduce r cl = + let cl = mk_clause cl in + Tactics.reduce r cl + let simpl flags where cl = let where = Option.map map_pattern_with_occs where in + let cl = mk_clause cl in Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> let flags = { flags with rConst } in Tactics.reduce (Simpl (flags, where)) cl let cbv flags cl = + let cl = mk_clause cl in Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> let flags = { flags with rConst } in Tactics.reduce (Cbv flags) cl let cbn flags cl = + let cl = mk_clause cl in Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> let flags = { flags with rConst } in Tactics.reduce (Cbn flags) cl let lazy_ flags cl = + let cl = mk_clause cl in Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> let flags = { flags with rConst } in Tactics.reduce (Lazy flags) cl let unfold occs cl = + let cl = mk_clause cl in let map (gr, occ) = + let occ = mk_occurrences_expr occ in get_evaluable_reference gr >>= fun gr -> Proofview.tclUNIT (occ, gr) in Proofview.Monad.List.map map occs >>= fun occs -> Tactics.reduce (Unfold occs) cl let pattern where cl = - let where = List.map (fun (c, occ) -> (occ, c)) where in + let where = List.map (fun (c, occ) -> (mk_occurrences_expr occ, c)) where in + let cl = mk_clause cl in Tactics.reduce (Pattern where) cl let vm where cl = let where = Option.map map_pattern_with_occs where in + let cl = mk_clause cl in Tactics.reduce (CbvVm where) cl let native where cl = let where = Option.map map_pattern_with_occs where in + let cl = mk_clause cl in Tactics.reduce (CbvNative where) cl let eval_fun red c = @@ -244,6 +286,7 @@ let eval_lazy flags c = let eval_unfold occs c = let map (gr, occ) = + let occ = mk_occurrences_expr occ in get_evaluable_reference gr >>= fun gr -> Proofview.tclUNIT (occ, gr) in Proofview.Monad.List.map map occs >>= fun occs -> @@ -253,7 +296,7 @@ let eval_fold cl c = eval_fun (Fold cl) c let eval_pattern where c = - let where = List.map (fun (pat, occ) -> (occ, pat)) where in + let where = List.map (fun (pat, occ) -> (mk_occurrences_expr occ, pat)) where in eval_fun (Pattern where) c let eval_vm where c = @@ -303,6 +346,7 @@ let injection ev ipat arg = let autorewrite ~all by ids cl = let conds = if all then Some Equality.AllMatches else None in let ids = List.map Id.to_string ids in + let cl = mk_clause cl in match by with | None -> Autorewrite.auto_multi_rewrite ?conds ids cl | Some by -> Autorewrite.auto_multi_rewrite_with ?conds by ids cl diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli index 842b09c22f..3d64e7ec8c 100644 --- a/src/tac2tactics.mli +++ b/src/tac2tactics.mli @@ -7,7 +7,6 @@ (************************************************************************) open Names -open Locus open Globnames open Tac2expr open EConstr @@ -31,6 +30,8 @@ val elim : evars_flag -> constr_with_bindings -> constr_with_bindings option -> val general_case_analysis : evars_flag -> constr_with_bindings -> unit tactic +val generalize : (constr * occurrences * Name.t) list -> unit tactic + val constructor_tac : evars_flag -> int option -> int -> bindings -> unit tactic val left_with_bindings : evars_flag -> bindings -> unit tactic @@ -40,6 +41,8 @@ val split_with_bindings : evars_flag -> bindings -> unit tactic val rewrite : evars_flag -> rewriting list -> clause -> unit tactic option -> unit tactic +val symmetry : clause -> unit tactic + val forward : bool -> unit tactic option option -> intro_pattern option -> constr -> unit tactic @@ -48,8 +51,10 @@ val assert_ : assertion -> unit tactic val letin_pat_tac : evars_flag -> (bool * intro_pattern_naming) option -> Name.t -> (Evd.evar_map * constr) -> clause -> unit tactic +val reduce : Redexpr.red_expr -> clause -> unit tactic + val simpl : global_reference glob_red_flag -> - (Pattern.constr_pattern * occurrences_expr) option -> clause -> unit tactic + (Pattern.constr_pattern * occurrences) option -> clause -> unit tactic val cbv : global_reference glob_red_flag -> clause -> unit tactic @@ -57,20 +62,20 @@ val cbn : global_reference glob_red_flag -> clause -> unit tactic val lazy_ : global_reference glob_red_flag -> clause -> unit tactic -val unfold : (global_reference * occurrences_expr) list -> clause -> unit tactic +val unfold : (global_reference * occurrences) list -> clause -> unit tactic -val pattern : (constr * occurrences_expr) list -> clause -> unit tactic +val pattern : (constr * occurrences) list -> clause -> unit tactic -val vm : (Pattern.constr_pattern * occurrences_expr) option -> clause -> unit tactic +val vm : (Pattern.constr_pattern * occurrences) option -> clause -> unit tactic -val native : (Pattern.constr_pattern * occurrences_expr) option -> clause -> unit tactic +val native : (Pattern.constr_pattern * occurrences) option -> clause -> unit tactic val eval_red : constr -> constr tactic val eval_hnf : constr -> constr tactic val eval_simpl : global_reference glob_red_flag -> - (Pattern.constr_pattern * occurrences_expr) option -> constr -> constr tactic + (Pattern.constr_pattern * occurrences) option -> constr -> constr tactic val eval_cbv : global_reference glob_red_flag -> constr -> constr tactic @@ -78,15 +83,15 @@ val eval_cbn : global_reference glob_red_flag -> constr -> constr tactic val eval_lazy : global_reference glob_red_flag -> constr -> constr tactic -val eval_unfold : (global_reference * occurrences_expr) list -> constr -> constr tactic +val eval_unfold : (global_reference * occurrences) list -> constr -> constr tactic val eval_fold : constr list -> constr -> constr tactic -val eval_pattern : (EConstr.t * occurrences_expr) list -> constr -> constr tactic +val eval_pattern : (EConstr.t * occurrences) list -> constr -> constr tactic -val eval_vm : (Pattern.constr_pattern * occurrences_expr) option -> constr -> constr tactic +val eval_vm : (Pattern.constr_pattern * occurrences) option -> constr -> constr tactic -val eval_native : (Pattern.constr_pattern * occurrences_expr) option -> constr -> constr tactic +val eval_native : (Pattern.constr_pattern * occurrences) option -> constr -> constr tactic val discriminate : evars_flag -> destruction_arg option -> unit tactic diff --git a/src/tac2types.mli b/src/tac2types.mli index 6845de8c7c..1cacbefc88 100644 --- a/src/tac2types.mli +++ b/src/tac2types.mli @@ -53,11 +53,26 @@ and or_and_intro_pattern = | IntroOrPattern of intro_pattern list list | IntroAndPattern of intro_pattern list +type occurrences = +| AllOccurrences +| AllOccurrencesBut of int list +| NoOccurrences +| OnlyOccurrences of int list + +type hyp_location_flag = Locus.hyp_location_flag = +| InHyp | InHypTypeOnly | InHypValueOnly + +type hyp_location = Id.t * occurrences * hyp_location_flag + +type clause = + { onhyps : hyp_location list option; + concl_occs : occurrences } + type induction_clause = destruction_arg * intro_pattern_naming option * or_and_intro_pattern option * - Locus.clause option + clause option type multi = Misctypes.multi = | Precisely of int -- cgit v1.2.3 From 0145084daa86b35a1d2a8285c4e16a9a231e3652 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 1 Oct 2017 19:10:11 +0200 Subject: Using Ltac2 native closures in some tactic APIs. --- src/tac2ffi.ml | 1 + src/tac2ffi.mli | 1 + src/tac2stdlib.ml | 4 ++-- src/tac2tactics.ml | 7 ++++++- src/tac2types.mli | 6 ++++-- 5 files changed, 14 insertions(+), 5 deletions(-) diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index 7960d4d45f..923a29e5c5 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -329,6 +329,7 @@ let reference = { type ('a, 'b) fun1 = closure let fun1 (r0 : 'a repr) (r1 : 'b repr) : ('a, 'b) fun1 repr = closure +let to_fun1 r0 r1 f = to_closure f let rec apply : type a. a arity -> a -> valexpr list -> valexpr Proofview.tactic = fun arity f args -> match args, arity with diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index 36743f3346..cb6d5a1e49 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -143,6 +143,7 @@ type ('a, 'b) fun1 val app_fun1 : ('a, 'b) fun1 -> 'a repr -> 'b repr -> 'a -> 'b Proofview.tactic +val to_fun1 : 'a repr -> 'b repr -> valexpr -> ('a, 'b) fun1 val fun1 : 'a repr -> 'b repr -> ('a, 'b) fun1 repr val valexpr : valexpr repr diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 4bcbe69b07..2828bbc53f 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -108,7 +108,7 @@ and to_intro_pattern_action = function let map ipat = to_intro_pattern ipat in IntroInjection (Value.to_list map inj) | ValBlk (2, [| c; ipat |]) -> - let c = thaw c >>= fun c -> return (Value.to_constr c) in + let c = Value.to_fun1 Value.unit Value.constr c in IntroApplyOn (c, to_intro_pattern ipat) | ValBlk (3, [| b |]) -> IntroRewrite (Value.to_bool b) | _ -> assert false @@ -143,7 +143,7 @@ let to_induction_clause v = match Value.to_tuple v with let to_assertion v = match Value.to_block v with | (0, [| ipat; t; tac |]) -> - let to_tac t = Proofview.tclIGNORE (thaw t) in + let to_tac t = Value.to_fun1 Value.unit Value.unit t in let ipat = Value.to_option to_intro_pattern ipat in let t = Value.to_constr t in let tac = Value.to_option to_tac tac in diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index b55bd5c1b8..0b25ebb378 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -17,6 +17,7 @@ open Proofview open Proofview.Notations let return = Proofview.tclUNIT +let thaw r f = Tac2ffi.app_fun1 f Tac2ffi.unit r () let tactic_infer_flags with_evar = { Pretyping.use_typeclasses = true; @@ -31,6 +32,9 @@ let delayed_of_tactic tac env sigma = let c, pv, _, _ = Proofview.apply env tac pv in (sigma, c) +let delayed_of_thunk r tac env sigma = + delayed_of_tactic (thaw r tac) env sigma + let mk_bindings = function | ImplicitBindings l -> Misctypes.ImplicitBindings l | ExplicitBindings l -> @@ -55,7 +59,7 @@ and mk_intro_pattern_action = function | IntroOrAndPattern ipat -> Misctypes.IntroOrAndPattern (mk_or_and_intro_pattern ipat) | IntroInjection ipats -> Misctypes.IntroInjection (List.map mk_intro_pattern ipats) | IntroApplyOn (c, ipat) -> - let c = Loc.tag @@ delayed_of_tactic c in + let c = Loc.tag @@ delayed_of_thunk Tac2ffi.constr c in Misctypes.IntroApplyOn (c, mk_intro_pattern ipat) | IntroRewrite b -> Misctypes.IntroRewrite b @@ -172,6 +176,7 @@ let assert_ = function Tactics.forward true None (Some ipat) c | AssertType (ipat, c, tac) -> let ipat = Option.map mk_intro_pattern ipat in + let tac = Option.map (fun tac -> thaw Tac2ffi.unit tac) tac in Tactics.forward true (Some tac) ipat c let letin_pat_tac ev ipat na c cl = diff --git a/src/tac2types.mli b/src/tac2types.mli index 1cacbefc88..a7b0ceed6e 100644 --- a/src/tac2types.mli +++ b/src/tac2types.mli @@ -15,6 +15,8 @@ open Proofview type evars_flag = bool type advanced_flag = bool +type 'a thunk = (unit, 'a) Tac2ffi.fun1 + type quantified_hypothesis = Misctypes.quantified_hypothesis = | AnonHyp of int | NamedHyp of Id.t @@ -47,7 +49,7 @@ and intro_pattern_action = | IntroWildcard | IntroOrAndPattern of or_and_intro_pattern | IntroInjection of intro_pattern list -| IntroApplyOn of EConstr.t tactic * intro_pattern +| IntroApplyOn of EConstr.t thunk * intro_pattern | IntroRewrite of bool and or_and_intro_pattern = | IntroOrPattern of intro_pattern list list @@ -86,5 +88,5 @@ type rewriting = constr_with_bindings tactic type assertion = -| AssertType of intro_pattern option * constr * unit tactic option +| AssertType of intro_pattern option * constr * unit thunk option | AssertValue of Id.t * constr -- cgit v1.2.3 From 78832226d6e472a6592dfffe91242613ec76841c Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 1 Oct 2017 20:33:26 +0200 Subject: Abstracting away the implementation of value representations. --- src/tac2core.ml | 6 +++--- src/tac2ffi.ml | 3 +++ src/tac2ffi.mli | 10 ++++------ src/tac2print.ml | 2 +- 4 files changed, 11 insertions(+), 10 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 9e3cefc6f5..1756e44f76 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -157,15 +157,15 @@ let define_primitive name arity f = let define0 name f = define_primitive name arity_one (fun _ -> f) let define1 name r0 f = define_primitive name arity_one begin fun x -> - f (r0.Value.r_to x) + f (Value.repr_to r0 x) end let define2 name r0 r1 f = define_primitive name (arity_suc arity_one) begin fun x y -> - f (r0.Value.r_to x) (r1.Value.r_to y) + f (Value.repr_to r0 x) (Value.repr_to r1 y) end let define3 name r0 r1 r2 f = define_primitive name (arity_suc (arity_suc arity_one)) begin fun x y z -> - f (r0.Value.r_to x) (r1.Value.r_to y) (r2.Value.r_to z) + f (Value.repr_to r0 x) (Value.repr_to r1 y) (Value.repr_to r2 z) end (** Printing *) diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index 923a29e5c5..c612cb85a5 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -76,6 +76,9 @@ type 'a repr = { r_id : bool; } +let repr_of r x = r.r_of x +let repr_to r x = r.r_to x + (** Dynamic tags *) let val_exn = Val.create "exn" diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index cb6d5a1e49..9f200b439d 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -49,12 +49,10 @@ end (** {5 Ltac2 FFI} *) -type 'a repr = { - r_of : 'a -> valexpr; - r_to : valexpr -> 'a; - r_id : bool; - (** True if the functions above are physical identities. *) -} +type 'a repr + +val repr_of : 'a repr -> 'a -> valexpr +val repr_to : 'a repr -> valexpr -> 'a (** These functions allow to convert back and forth between OCaml and Ltac2 data representation. The [to_*] functions raise an anomaly whenever the data diff --git a/src/tac2print.ml b/src/tac2print.ml index 45360a61f4..018346ad6a 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -401,7 +401,7 @@ let rec pr_valexpr env sigma v t = match kind t with end | GTypArrow _ -> str "" | GTypRef (Tuple _, tl) -> - let blk = Array.to_list (snd (block.r_to v)) in + let blk = Array.to_list (snd (to_block v)) in if List.length blk == List.length tl then let prs = List.map2 (fun v t -> pr_valexpr env sigma v t) blk tl in hv 2 (str "(" ++ prlist_with_sep pr_comma (fun p -> p) prs ++ str ")") -- cgit v1.2.3 From f0fb8686377815978556e7d554f27de6406d1fed Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 7 Oct 2017 22:46:45 +0200 Subject: Fix coq/ltac2#30: Compilation broken since -safe-string PR got merged. --- src/tac2core.ml | 5 +++-- src/tac2ffi.mli | 6 +++--- src/tac2print.ml | 2 +- 3 files changed, 7 insertions(+), 6 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 1756e44f76..1489c50c13 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -234,11 +234,11 @@ let () = define2 "ident_equal" ident ident begin fun id1 id2 -> end let () = define1 "ident_to_string" ident begin fun id -> - return (Value.of_string (Id.to_string id)) + return (Value.of_string (Bytes.of_string (Id.to_string id))) end let () = define1 "ident_of_string" string begin fun s -> - let id = try Some (Id.of_string s) with _ -> None in + let id = try Some (Id.of_string (Bytes.to_string s)) with _ -> None in return (Value.of_option Value.of_ident id) end @@ -714,6 +714,7 @@ let () = define2 "abstract" (option ident) closure begin fun id f -> end let () = define2 "time" (option string) closure begin fun s f -> + let s = Option.map Bytes.to_string s in Proofview.tclTIME s (thaw f) end diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index 9f200b439d..9c3b512036 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -74,9 +74,9 @@ val of_char : char -> valexpr val to_char : valexpr -> char val char : char repr -val of_string : string -> valexpr -val to_string : valexpr -> string -val string : string repr +val of_string : Bytes.t -> valexpr +val to_string : valexpr -> Bytes.t +val string : Bytes.t repr val of_list : ('a -> valexpr) -> 'a list -> valexpr val to_list : (valexpr -> 'a) -> valexpr -> 'a list diff --git a/src/tac2print.ml b/src/tac2print.ml index 018346ad6a..5f34b54ee6 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -426,7 +426,7 @@ end let () = register_init "string" begin fun s -> let s = to_string s in - Pp.quote (Pp.str s) + Pp.quote (Pp.str (Bytes.to_string s)) end let () = register_init "ident" begin fun id -> -- cgit v1.2.3 From 533b5ee5a3c5dd4c2e54d85dba9485722bb21db1 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 7 Oct 2017 22:47:48 +0200 Subject: Remove unused warnings. --- src/tac2core.ml | 4 ---- src/tac2entries.ml | 1 - src/tac2env.ml | 2 -- src/tac2env.mli | 1 - src/tac2expr.mli | 1 - src/tac2ffi.ml | 1 - src/tac2intern.ml | 6 ------ src/tac2interp.ml | 1 - src/tac2print.ml | 1 - src/tac2quote.ml | 2 -- src/tac2quote.mli | 1 - src/tac2stdlib.ml | 2 -- src/tac2tactics.ml | 2 -- 13 files changed, 25 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 1489c50c13..0304286639 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -6,13 +6,11 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open CSig open Util open Pp open Names open Genarg open Tac2env -open Tac2dyn open Tac2expr open Tac2entries.Pltac open Proofview.Notations @@ -105,7 +103,6 @@ let err_matchfailure = let thaw f = Tac2ffi.apply f [v_unit] let fatal_flag : unit Exninfo.t = Exninfo.make () -let fatal_info = Exninfo.add Exninfo.null fatal_flag () let set_bt info = if !Tac2interp.print_ltac2_backtrace then @@ -397,7 +394,6 @@ let () = define1 "constr_kind" constr begin fun c -> end let () = define1 "constr_make" valexpr begin fun knd -> - let open Constr in let c = match Tac2ffi.to_block knd with | (0, [|n|]) -> let n = Value.to_int n in diff --git a/src/tac2entries.ml b/src/tac2entries.ml index cd4b701ca7..f24c409ad7 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -16,7 +16,6 @@ open Nametab open Tac2expr open Tac2print open Tac2intern -open Vernacexpr (** Grammar entries *) diff --git a/src/tac2env.ml b/src/tac2env.ml index 0aa2da77ae..2f1124c156 100644 --- a/src/tac2env.ml +++ b/src/tac2env.ml @@ -6,11 +6,9 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open CErrors open Util open Names open Libnames -open Tac2dyn open Tac2expr open Tac2ffi diff --git a/src/tac2env.mli b/src/tac2env.mli index b82923765d..022c518143 100644 --- a/src/tac2env.mli +++ b/src/tac2env.mli @@ -10,7 +10,6 @@ open Genarg open Names open Libnames open Nametab -open Tac2dyn open Tac2expr open Tac2ffi diff --git a/src/tac2expr.mli b/src/tac2expr.mli index c787870c65..e57b0ba3ef 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -7,7 +7,6 @@ (************************************************************************) open Loc -open Genarg open Names open Libnames diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index c612cb85a5..19d4259b55 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -9,7 +9,6 @@ open Util open Names open Globnames -open Genarg open Tac2dyn open Tac2expr open Proofview.Notations diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 0efd9a3005..7b35cd55aa 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -8,7 +8,6 @@ open Pp open Util -open Genarg open CErrors open Names open Libnames @@ -23,13 +22,8 @@ let coq_type n = KerName.make2 Tac2env.coq_prefix (Label.make n) let t_int = coq_type "int" let t_string = coq_type "string" -let t_array = coq_type "array" -let t_list = coq_type "list" let t_constr = coq_type "constr" -let c_nil = GTacCst (Other t_list, 0, []) -let c_cons e el = GTacCst (Other t_list, 0, [e; el]) - (** Union find *) module UF : diff --git a/src/tac2interp.ml b/src/tac2interp.ml index db30f52772..6f158ac66e 100644 --- a/src/tac2interp.ml +++ b/src/tac2interp.ml @@ -9,7 +9,6 @@ open Util open Pp open CErrors -open Genarg open Names open Proofview.Notations open Tac2expr diff --git a/src/tac2print.ml b/src/tac2print.ml index 5f34b54ee6..8e4947e332 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -8,7 +8,6 @@ open Util open Pp -open Genarg open Names open Tac2expr open Tac2env diff --git a/src/tac2quote.ml b/src/tac2quote.ml index e89f37f2ba..e967067161 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -40,8 +40,6 @@ let pattern_core n = kername pattern_prefix n let global_ref ?loc kn = Loc.tag ?loc @@ CTacRef (AbsKn (TacConstant kn)) -let dummy_loc = Loc.make_loc (-1, -1) - let constructor ?loc kn args = let cst = Loc.tag ?loc @@ CTacCst (AbsKn (Other kn)) in if List.is_empty args then cst diff --git a/src/tac2quote.mli b/src/tac2quote.mli index b9cae23e63..148e6818dd 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Util open Loc open Names open Tac2dyn diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 2828bbc53f..99fa0370e1 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -7,12 +7,10 @@ (************************************************************************) open Names -open Globnames open Genredexpr open Tac2expr open Tac2ffi open Tac2types -open Tac2tactics open Proofview.Notations module Value = Tac2ffi diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index 0b25ebb378..b496f5046f 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -10,10 +10,8 @@ open Pp open Util open Names open Globnames -open Misctypes open Tac2types open Genredexpr -open Proofview open Proofview.Notations let return = Proofview.tclUNIT -- cgit v1.2.3 From 0b26bfc8e068e1e95eeea9db0c3bda7436ac8338 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 29 Sep 2017 17:57:25 +0200 Subject: Implementing the Constr.in_context function. --- src/tac2core.ml | 32 ++++++++++++++++++++++++++++++++ theories/Constr.v | 5 +++++ 2 files changed, 37 insertions(+) diff --git a/src/tac2core.ml b/src/tac2core.ml index 0304286639..468fddaddf 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -494,6 +494,38 @@ let () = define3 "constr_closenl" (list ident) int constr begin fun ids k c -> return (Value.of_constr ans) end +let () = define3 "constr_in_context" ident constr closure begin fun id t c -> + Proofview.Goal.goals >>= function + | [gl] -> + gl >>= fun gl -> + let env = Proofview.Goal.env gl in + let sigma = Proofview.Goal.sigma gl in + let has_var = + try + let _ = Environ.lookup_named_val id (Environ.named_context_val env) in + true + with Not_found -> false + in + if has_var then + Tacticals.New.tclZEROMSG (str "Variable already exists") + else + let open Context.Named.Declaration in + let nenv = EConstr.push_named (LocalAssum (id, t)) env in + let (sigma, (evt, _)) = Evarutil.new_type_evar nenv sigma Evd.univ_flexible in + let (sigma, evk) = Evarutil.new_pure_evar (Environ.named_context_val nenv) sigma evt in + Proofview.Unsafe.tclEVARS sigma >>= fun () -> + Proofview.Unsafe.tclSETGOALS [evk] >>= fun () -> + thaw c >>= fun _ -> + Proofview.Unsafe.tclSETGOALS [Proofview.Goal.goal (Proofview.Goal.assume gl)] >>= fun () -> + let args = List.map (fun d -> EConstr.mkVar (get_id d)) (EConstr.named_context env) in + let args = Array.of_list (EConstr.mkRel 1 :: args) in + let ans = EConstr.mkEvar (evk, args) in + let ans = EConstr.mkLambda (Name id, t, ans) in + return (Value.of_constr ans) + | _ -> + throw err_notfocussed +end + (** Patterns *) let () = define2 "pattern_matches" pattern constr begin fun pat c -> diff --git a/theories/Constr.v b/theories/Constr.v index 3e67a486cf..072c613920 100644 --- a/theories/Constr.v +++ b/theories/Constr.v @@ -58,3 +58,8 @@ Ltac2 @ external closenl : ident list -> int -> constr -> constr := "ltac2" "con [Rel(k); ...; Rel(k+n-1)] in [c]. If two names are identical, the one of least index is kept. *) End Unsafe. + +Ltac2 @ external in_context : ident -> constr -> (unit -> unit) -> constr := "ltac2" "constr_in_context". +(** On a focussed goal [Γ ⊢ A], [in_context id c tac] evaluates [tac] in a + focussed goal [Γ, id : c ⊢ ?X] and returns [fun (id : c) => t] where [t] is + the proof built by the tactic. *) -- cgit v1.2.3 From d4172d2c7a48d932b42248fe57c6c2a87ac57e30 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 25 Oct 2017 12:13:27 +0200 Subject: Stubs for goal matching: quotation and matching function. --- _CoqProject | 2 + src/g_ltac2.ml4 | 33 +++++-- src/ltac2_plugin.mlpack | 1 + src/tac2core.ml | 26 ++++++ src/tac2entries.ml | 1 + src/tac2entries.mli | 1 + src/tac2ffi.ml | 10 +++ src/tac2ffi.mli | 2 + src/tac2match.ml | 232 ++++++++++++++++++++++++++++++++++++++++++++++++ src/tac2match.mli | 33 +++++++ src/tac2qexpr.mli | 21 ++++- src/tac2quote.ml | 150 ++++++++++++++++++++++--------- src/tac2quote.mli | 2 + theories/Notations.v | 60 ++++++------- theories/Pattern.v | 29 +++++- 15 files changed, 515 insertions(+), 88 deletions(-) create mode 100644 src/tac2match.ml create mode 100644 src/tac2match.mli diff --git a/_CoqProject b/_CoqProject index 43e9b76991..df8cb00b76 100644 --- a/_CoqProject +++ b/_CoqProject @@ -21,6 +21,8 @@ src/tac2ffi.mli src/tac2qexpr.mli src/tac2quote.ml src/tac2quote.mli +src/tac2match.ml +src/tac2match.mli src/tac2core.ml src/tac2core.mli src/tac2tactics.ml diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index c92a242637..be6c2de5a9 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -376,7 +376,7 @@ GEXTEND Gram GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause q_rewriting q_clause q_dispatch q_occurrences q_strategy_flag q_destruction_arg q_reference q_with_bindings q_constr_matching - q_hintdb q_move_location q_pose q_assert; + q_goal_matching q_hintdb q_move_location q_pose q_assert; anti: [ [ "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] ; @@ -682,14 +682,12 @@ GEXTEND Gram ; match_pattern: [ [ IDENT "context"; id = OPT Prim.ident; - "["; pat = Constr.lconstr_pattern; "]" -> (Some id, pat) - | pat = Constr.lconstr_pattern -> (None, pat) ] ] + "["; pat = Constr.lconstr_pattern; "]" -> Loc.tag ~loc:!@loc @@ QConstrMatchContext (id, pat) + | pat = Constr.lconstr_pattern -> Loc.tag ~loc:!@loc @@ QConstrMatchPattern pat ] ] ; match_rule: [ [ mp = match_pattern; "=>"; tac = tac2expr -> - match mp with - | None, pat -> Loc.tag ~loc:!@loc @@ QConstrMatchPattern (pat, tac) - | Some oid, pat -> Loc.tag ~loc:!@loc @@ QConstrMatchContext (oid, pat, tac) + Loc.tag ~loc:!@loc @@ (mp, tac) ] ] ; match_list: @@ -699,6 +697,29 @@ GEXTEND Gram q_constr_matching: [ [ m = match_list -> m ] ] ; + gmatch_hyp_pattern: + [ [ na = Prim.name; ":"; pat = match_pattern -> (na, pat) ] ] + ; + gmatch_pattern: + [ [ "["; hl = LIST0 gmatch_hyp_pattern SEP ","; "|-"; p = match_pattern; "]" -> + Loc.tag ~loc:!@loc @@ { + q_goal_match_concl = p; + q_goal_match_hyps = hl; + } + ] ] + ; + gmatch_rule: + [ [ mp = gmatch_pattern; "=>"; tac = tac2expr -> + Loc.tag ~loc:!@loc @@ (mp, tac) + ] ] + ; + gmatch_list: + [ [ mrl = LIST1 gmatch_rule SEP "|" -> Loc.tag ~loc:!@loc @@ mrl + | "|"; mrl = LIST1 gmatch_rule SEP "|" -> Loc.tag ~loc:!@loc @@ mrl ] ] + ; + q_goal_matching: + [ [ m = gmatch_list -> m ] ] + ; move_location: [ [ "at"; IDENT "top" -> Loc.tag ~loc:!@loc @@ QMoveFirst | "at"; IDENT "bottom" -> Loc.tag ~loc:!@loc @@ QMoveLast diff --git a/src/ltac2_plugin.mlpack b/src/ltac2_plugin.mlpack index a2237f4d26..40b91e4b53 100644 --- a/src/ltac2_plugin.mlpack +++ b/src/ltac2_plugin.mlpack @@ -6,6 +6,7 @@ Tac2intern Tac2interp Tac2entries Tac2quote +Tac2match Tac2core Tac2tactics Tac2stdlib diff --git a/src/tac2core.ml b/src/tac2core.ml index 468fddaddf..c4502f8eae 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -528,6 +528,12 @@ end (** Patterns *) +let empty_context = EConstr.mkMeta Constr_matching.special_meta + +let () = define0 "pattern_empty_context" begin + return (Value.of_constr empty_context) +end + let () = define2 "pattern_matches" pattern constr begin fun pat c -> pf_apply begin fun env sigma -> let ans = @@ -592,6 +598,25 @@ let () = define2 "pattern_matches_subterm_vect" pattern constr begin fun pat c - end end +let () = define3 "pattern_matches_goal" bool (list (pair bool pattern)) (pair bool pattern) begin fun rev hp cp -> + assert_focussed >>= fun () -> + Proofview.Goal.enter_one begin fun gl -> + let env = Proofview.Goal.env gl in + let sigma = Proofview.Goal.sigma gl in + let concl = Proofview.Goal.concl gl in + let mk_pattern (b, pat) = if b then Tac2match.MatchContext pat else Tac2match.MatchPattern pat in + let r = (List.map mk_pattern hp, mk_pattern cp) in + Tac2match.match_goal env sigma concl ~rev r >>= fun (hyps, ctx, subst) -> + let of_ctxopt ctx = Value.of_constr (Option.default empty_context ctx) in + let hids = Value.of_array Value.of_ident (Array.map_of_list fst hyps) in + let hctx = Value.of_array of_ctxopt (Array.map_of_list snd hyps) in + let subs = Value.of_array Value.of_constr (Array.map_of_list snd (Id.Map.bindings subst)) in + let cctx = of_ctxopt ctx in + let ans = Value.of_tuple [| hids; hctx; subs; cctx |] in + Proofview.tclUNIT ans + end +end + let () = define2 "pattern_instantiate" constr constr begin fun ctx c -> let ctx = EConstr.Unsafe.to_constr ctx in let c = EConstr.Unsafe.to_constr c in @@ -1146,6 +1171,7 @@ let () = add_expr_scope "move_location" q_move_location Tac2quote.of_move_locati let () = add_expr_scope "pose" q_pose Tac2quote.of_pose let () = add_expr_scope "assert" q_assert Tac2quote.of_assertion let () = add_expr_scope "constr_matching" q_constr_matching Tac2quote.of_constr_matching +let () = add_expr_scope "goal_matching" q_goal_matching Tac2quote.of_goal_matching let () = add_generic_scope "constr" Pcoq.Constr.constr Tac2quote.wit_constr let () = add_generic_scope "open_constr" Pcoq.Constr.constr Tac2quote.wit_open_constr diff --git a/src/tac2entries.ml b/src/tac2entries.ml index f24c409ad7..04bf21f656 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -37,6 +37,7 @@ let q_occurrences = Pcoq.Gram.entry_create "tactic:q_occurrences" let q_reference = Pcoq.Gram.entry_create "tactic:q_reference" let q_strategy_flag = Pcoq.Gram.entry_create "tactic:q_strategy_flag" let q_constr_matching = Pcoq.Gram.entry_create "tactic:q_constr_matching" +let q_goal_matching = Pcoq.Gram.entry_create "tactic:q_goal_matching" let q_hintdb = Pcoq.Gram.entry_create "tactic:q_hintdb" let q_move_location = Pcoq.Gram.entry_create "tactic:q_move_location" let q_pose = Pcoq.Gram.entry_create "tactic:q_pose" diff --git a/src/tac2entries.mli b/src/tac2entries.mli index 7bd512651c..b2a2dd4846 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -77,6 +77,7 @@ val q_occurrences : occurrences Pcoq.Gram.entry val q_reference : Libnames.reference or_anti Pcoq.Gram.entry val q_strategy_flag : strategy_flag Pcoq.Gram.entry val q_constr_matching : constr_matching Pcoq.Gram.entry +val q_goal_matching : goal_matching Pcoq.Gram.entry val q_hintdb : hintdb Pcoq.Gram.entry val q_move_location : move_location Pcoq.Gram.entry val q_pose : pose Pcoq.Gram.entry diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index 19d4259b55..b0c56cdf45 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -272,6 +272,16 @@ let to_tuple = function | ValBlk (0, cl) -> cl | _ -> assert false +let of_pair f g (x, y) = ValBlk (0, [|f x; g y|]) +let to_pair f g = function +| ValBlk (0, [|x; y|]) -> (f x, g y) +| _ -> assert false +let pair r0 r1 = { + r_of = (fun p -> of_pair r0.r_of r1.r_of p); + r_to = (fun p -> to_pair r0.r_to r1.r_to p); + r_id = false; +} + let of_array f vl = ValBlk (0, Array.map f vl) let to_array f = function | ValBlk (0, vl) -> Array.map f vl diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index 9c3b512036..6d4867a453 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -109,6 +109,8 @@ val array : 'a repr -> 'a array repr val of_tuple : valexpr array -> valexpr val to_tuple : valexpr -> valexpr array +val pair : 'a repr -> 'b repr -> ('a * 'b) repr + val of_option : ('a -> valexpr) -> 'a option -> valexpr val to_option : (valexpr -> 'a) -> valexpr -> 'a option val option : 'a repr -> 'a option repr diff --git a/src/tac2match.ml b/src/tac2match.ml new file mode 100644 index 0000000000..7a22e91b6b --- /dev/null +++ b/src/tac2match.ml @@ -0,0 +1,232 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* None + | None, Some c | Some c, None -> Some c + | Some c1, Some c2 -> + if equal_instances env sigma c1 c2 then Some c1 + else raise Not_coherent_metas + in + Id.Map.merge merge s1 s2 + +let matching_error = + CErrors.UserError (Some "tactic matching" , Pp.str "No matching clauses for match.") + +let imatching_error = (matching_error, Exninfo.null) + +(** A functor is introduced to share the environment and the + evar_map. They do not change and it would be a pity to introduce + closures everywhere just for the occasional calls to + {!equal_instances}. *) +module type StaticEnvironment = sig + val env : Environ.env + val sigma : Evd.evar_map +end +module PatternMatching (E:StaticEnvironment) = struct + + + (** {6 The pattern-matching monad } *) + + + (** To focus on the algorithmic portion of pattern-matching, the + bookkeeping is relegated to a monad: the composition of the + bactracking monad of {!IStream.t} with a "writer" effect. *) + (* spiwack: as we don't benefit from the various stream optimisations + of Haskell, it may be costly to give the monad in direct style such as + here. We may want to use some continuation passing style. *) + type 'a tac = 'a Proofview.tactic + type 'a m = { stream : 'r. ('a -> result -> 'r tac) -> result -> 'r tac } + + (** The empty substitution. *) + let empty_subst = Id.Map.empty + + (** Composes two substitutions using {!verify_metas_coherence}. It + must be a monoid with neutral element {!empty_subst}. Raises + [Not_coherent_metas] when composition cannot be achieved. *) + let subst_prod s1 s2 = + if is_empty_subst s1 then s2 + else if is_empty_subst s2 then s1 + else verify_metas_coherence E.env E.sigma s1 s2 + + (** Merge two writers (and ignore the first value component). *) + let merge m1 m2 = + try Some { + subst = subst_prod m1.subst m2.subst; + } + with Not_coherent_metas -> None + + (** Monadic [return]: returns a single success with empty substitutions. *) + let return (type a) (lhs:a) : a m = + { stream = fun k ctx -> k lhs ctx } + + (** Monadic bind: each success of [x] is replaced by the successes + of [f x]. The substitutions of [x] and [f x] are composed, + dropping the apparent successes when the substitutions are not + coherent. *) + let (>>=) (type a) (type b) (m:a m) (f:a -> b m) : b m = + { stream = fun k ctx -> m.stream (fun x ctx -> (f x).stream k ctx) ctx } + + (** A variant of [(>>=)] when the first argument returns [unit]. *) + let (<*>) (type a) (m:unit m) (y:a m) : a m = + { stream = fun k ctx -> m.stream (fun () ctx -> y.stream k ctx) ctx } + + (** Failure of the pattern-matching monad: no success. *) + let fail (type a) : a m = { stream = fun _ _ -> Proofview.tclZERO matching_error } + + let run (m : 'a m) = + let ctx = { + subst = empty_subst ; + } in + let eval x ctx = Proofview.tclUNIT (x, ctx) in + m.stream eval ctx + + (** Chooses in a list, in the same order as the list *) + let rec pick (l:'a list) (e, info) : 'a m = match l with + | [] -> { stream = fun _ _ -> Proofview.tclZERO ~info e } + | x :: l -> + { stream = fun k ctx -> Proofview.tclOR (k x ctx) (fun e -> (pick l e).stream k ctx) } + + let pick l = pick l imatching_error + + let put_subst subst : unit m = + let s = { subst } in + { stream = fun k ctx -> match merge s ctx with None -> Proofview.tclZERO matching_error | Some s -> k () s } + + (** {6 Pattern-matching} *) + + let pattern_match_term pat term = + match pat with + | MatchPattern p -> + begin + try + put_subst (Constr_matching.matches E.env E.sigma p term) <*> + return None + with Constr_matching.PatternMatchingFailure -> fail + end + | MatchContext p -> + + let rec map s (e, info) = + { stream = fun k ctx -> match IStream.peek s with + | IStream.Nil -> Proofview.tclZERO ~info e + | IStream.Cons ({ Constr_matching.m_sub = (_, subst); m_ctx }, s) -> + let nctx = { subst } in + match merge ctx nctx with + | None -> (map s (e, info)).stream k ctx + | Some nctx -> Proofview.tclOR (k (Some m_ctx) nctx) (fun e -> (map s e).stream k ctx) + } + in + map (Constr_matching.match_appsubterm E.env E.sigma p term) imatching_error + + let hyp_match_type pat hyps = + pick hyps >>= fun decl -> + let id = NamedDecl.get_id decl in + pattern_match_term pat (NamedDecl.get_type decl) >>= fun ctx -> + return (id, ctx) + + let hyp_match_body_and_type bodypat typepat hyps = + pick hyps >>= function + | LocalDef (id,body,hyp) -> + pattern_match_term bodypat body >>= fun ctx_body -> + pattern_match_term typepat hyp >>= fun ctx_typ -> + return (id, ctx_body, ctx_typ) + | LocalAssum (id,hyp) -> fail + + let hyp_match pat hyps = + match pat with + | typepat -> + hyp_match_type typepat hyps +(* | Def ((_,hypname),bodypat,typepat) -> *) +(* hyp_match_body_and_type hypname bodypat typepat hyps *) + + (** [hyp_pattern_list_match pats hyps lhs], matches the list of + patterns [pats] against the hypotheses in [hyps], and eventually + returns [lhs]. *) + let rec hyp_pattern_list_match pats hyps accu = + match pats with + | pat::pats -> + hyp_match pat hyps >>= fun (matched_hyp, hyp_ctx) -> + let select_matched_hyp decl = Id.equal (NamedDecl.get_id decl) matched_hyp in + let hyps = CList.remove_first select_matched_hyp hyps in + hyp_pattern_list_match pats hyps ((matched_hyp, hyp_ctx) :: accu) + | [] -> return accu + + let rule_match_goal hyps concl = function + | (hyppats,conclpat) -> + (* the rules are applied from the topmost one (in the concrete + syntax) to the bottommost. *) + let hyppats = List.rev hyppats in + pattern_match_term conclpat concl >>= fun ctx_concl -> + hyp_pattern_list_match hyppats hyps [] >>= fun hyps -> + return (hyps, ctx_concl) + +end + +let match_goal env sigma concl ~rev rule = + let open Proofview.Notations in + let hyps = EConstr.named_context env in + let hyps = if rev then List.rev hyps else hyps in + let module E = struct + let env = env + let sigma = sigma + end in + let module M = PatternMatching(E) in + M.run (M.rule_match_goal hyps concl rule) >>= fun ((hyps, ctx_concl), subst) -> + Proofview.tclUNIT (hyps, ctx_concl, subst.subst) diff --git a/src/tac2match.mli b/src/tac2match.mli new file mode 100644 index 0000000000..cf64542f40 --- /dev/null +++ b/src/tac2match.mli @@ -0,0 +1,33 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* + Evd.evar_map -> + constr -> + rev:bool -> + match_rule -> + ((Id.t * context option) list * (** List of hypotheses matching: name + context *) + context option * (** Context for conclusion *) + Pattern.patvar_map (** Pattern variable substitution *)) Proofview.tactic diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index cb43a980de..229cece7c4 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -118,14 +118,27 @@ type red_flag = red_flag_r located type strategy_flag = red_flag list located -type constr_match_branch_r = -| QConstrMatchPattern of Constrexpr.constr_expr * raw_tacexpr -| QConstrMatchContext of Id.t option * Constrexpr.constr_expr * raw_tacexpr +type constr_match_pattern_r = +| QConstrMatchPattern of Constrexpr.constr_expr +| QConstrMatchContext of Id.t option * Constrexpr.constr_expr -type constr_match_branch = constr_match_branch_r located +type constr_match_pattern = constr_match_pattern_r located + +type constr_match_branch = (constr_match_pattern * raw_tacexpr) located type constr_matching = constr_match_branch list located +type goal_match_pattern_r = { + q_goal_match_concl : constr_match_pattern; + q_goal_match_hyps : (Name.t located * constr_match_pattern) list; +} + +type goal_match_pattern = goal_match_pattern_r located + +type goal_match_branch = (goal_match_pattern * raw_tacexpr) located + +type goal_matching = goal_match_branch list located + type hintdb_r = | QHintAll | QHintDbs of Id.t located or_anti list diff --git a/src/tac2quote.ml b/src/tac2quote.ml index e967067161..9728c3af93 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -307,62 +307,126 @@ let of_hintdb (loc, hdb) = match hdb with | QHintAll -> of_option ?loc (fun l -> of_list (fun id -> of_anti of_ident id) l) None | QHintDbs ids -> of_option ?loc (fun l -> of_list (fun id -> of_anti of_ident id) l) (Some ids) +let check_pattern_id ?loc id = + if Tac2env.is_constructor (Libnames.qualid_of_ident id) then + CErrors.user_err ?loc (str "Invalid pattern binding name " ++ Id.print id) + let pattern_vars pat = let rec aux () accu pat = match pat.CAst.v with - | Constrexpr.CPatVar id -> Id.Set.add id accu - | Constrexpr.CEvar (id, []) -> Id.Set.add id accu + | Constrexpr.CPatVar id + | Constrexpr.CEvar (id, []) -> + let () = check_pattern_id ?loc:pat.CAst.loc id in + Id.Set.add id accu | _ -> Topconstr.fold_constr_expr_with_binders (fun _ () -> ()) aux () accu pat in aux () Id.Set.empty pat -let of_constr_matching (loc, m) = - let check_id loc id = - if Tac2env.is_constructor (Libnames.qualid_of_ident id) then - CErrors.user_err ?loc (str "Invalid pattern binding name " ++ Id.print id) - in - let abstract_vars loc pat tac = - let vars = pattern_vars pat in - let na, tac = - if Id.Set.is_empty vars then (Anonymous, tac) - else - (** Trick: in order not to shadow a variable nor to choose an arbitrary - name, we reuse one which is going to be shadowed by the matched - variables anyways. *) - let id0 = Id.Set.choose vars in - let build_bindings id (n, accu) = - let () = check_id loc id in - let get = global_ref ?loc (kername array_prefix "get") in - let args = [of_variable (loc, id0); of_int (loc, n)] in - let e = Loc.tag ?loc @@ CTacApp (get, args) in - let accu = (Loc.tag ?loc @@ CPatVar (Name id), None, e) :: accu in - (n + 1, accu) - in - let (_, bnd) = Id.Set.fold build_bindings vars (0, []) in - let tac = Loc.tag ?loc @@ CTacLet (false, bnd, tac) in - (Name id0, tac) - in - Loc.tag ?loc @@ CTacFun ([Loc.tag ?loc @@ CPatVar na, None], tac) +let abstract_vars loc vars tac = + let get_name = function Name id -> Some id | Anonymous -> None in + let def = try Some (List.find_map get_name vars) with Not_found -> None in + let na, tac = match def with + | None -> (Anonymous, tac) + | Some id0 -> + (** Trick: in order not to shadow a variable nor to choose an arbitrary + name, we reuse one which is going to be shadowed by the matched + variables anyways. *) + let build_bindings (n, accu) na = match na with + | Anonymous -> (n + 1, accu) + | Name _ -> + let get = global_ref ?loc (kername array_prefix "get") in + let args = [of_variable (loc, id0); of_int (loc, n)] in + let e = Loc.tag ?loc @@ CTacApp (get, args) in + let accu = (Loc.tag ?loc @@ CPatVar na, None, e) :: accu in + (n + 1, accu) + in + let (_, bnd) = List.fold_left build_bindings (0, []) vars in + let tac = Loc.tag ?loc @@ CTacLet (false, bnd, tac) in + (Name id0, tac) in - let map (loc, p) = match p with - | QConstrMatchPattern (pat, tac) -> - let e = abstract_vars loc pat tac in - let pat = inj_wit ?loc wit_pattern pat in - constructor ?loc (pattern_core "ConstrMatchPattern") [pat; e] - | QConstrMatchContext (id, pat, tac) -> - let e = abstract_vars loc pat tac in - let na = match id with - | None -> Anonymous - | Some id -> - let () = check_id loc id in - Name id + Loc.tag ?loc @@ CTacFun ([Loc.tag ?loc @@ CPatVar na, None], tac) + +let extract_name ?loc oid = match oid with +| None -> Anonymous +| Some id -> + let () = check_pattern_id ?loc id in + Name id + +(** For every branch in the matching, generate a corresponding term of type + [(match_kind * pattern * (context -> constr array -> 'a))] + where the function binds the names from the pattern to the contents of the + constr array. *) +let of_constr_matching (loc, m) = + let map (loc, ((ploc, pat), tac)) = + let (knd, pat, na) = match pat with + | QConstrMatchPattern pat -> + let knd = constructor ?loc (pattern_core "MatchPattern") [] in + (knd, pat, Anonymous) + | QConstrMatchContext (id, pat) -> + let na = extract_name ?loc id in + let knd = constructor ?loc (pattern_core "MatchContext") [] in + (knd, pat, na) in + let vars = pattern_vars pat in + (** Order of elements is crucial here! *) + let vars = Id.Set.elements vars in + let vars = List.map (fun id -> Name id) vars in + let e = abstract_vars loc vars tac in let e = Loc.tag ?loc @@ CTacFun ([Loc.tag ?loc @@ CPatVar na, None], e) in - let pat = inj_wit ?loc wit_pattern pat in - constructor ?loc (pattern_core "ConstrMatchContext") [pat; e] + let pat = inj_wit ?loc:ploc wit_pattern pat in + of_tuple [knd; pat; e] in of_list ?loc map m +let of_pattern p = + inj_wit ?loc:p.CAst.loc wit_pattern p + +(** From the patterns and the body of the branch, generate: + - a goal pattern: (constr_match list * constr_match) + - a branch function (ident array -> context array -> constr array -> context -> 'a) +*) +let of_goal_matching (loc, gm) = + let mk_pat (loc, p) = match p with + | QConstrMatchPattern pat -> + let knd = constructor ?loc (pattern_core "MatchPattern") [] in + (Anonymous, pat, knd) + | QConstrMatchContext (id, pat) -> + let na = extract_name ?loc id in + let knd = constructor ?loc (pattern_core "MatchContext") [] in + (na, pat, knd) + in + let mk_gpat (loc, p) = + let concl_pat = p.q_goal_match_concl in + let hyps_pats = p.q_goal_match_hyps in + let (concl_ctx, concl_pat, concl_knd) = mk_pat concl_pat in + let vars = pattern_vars concl_pat in + let map accu (na, pat) = + let (ctx, pat, knd) = mk_pat pat in + let vars = pattern_vars pat in + (Id.Set.union vars accu, (na, ctx, pat, knd)) + in + let (vars, hyps_pats) = List.fold_left_map map vars hyps_pats in + let map (_, _, pat, knd) = of_tuple [knd; of_pattern pat] in + let concl = of_tuple [concl_knd; of_pattern concl_pat] in + let r = of_tuple [of_list ?loc map hyps_pats; concl] in + let hyps = List.map (fun ((_, na), _, _, _) -> na) hyps_pats in + let map (_, na, _, _) = na in + let hctx = List.map map hyps_pats in + (** Order of elements is crucial here! *) + let vars = Id.Set.elements vars in + let subst = List.map (fun id -> Name id) vars in + (r, hyps, hctx, subst, concl_ctx) + in + let map (loc, (pat, tac)) = + let (pat, hyps, hctx, subst, cctx) = mk_gpat pat in + let tac = abstract_vars loc hyps tac in + let tac = abstract_vars loc hctx tac in + let tac = abstract_vars loc subst tac in + let tac = abstract_vars loc [cctx] tac in + of_tuple ?loc [pat; tac] + in + of_list ?loc map gm + let of_move_location (loc, mv) = match mv with | QMoveAfter id -> std_constructor ?loc "MoveAfter" [of_anti of_ident id] | QMoveBefore id -> std_constructor ?loc "MoveBefore" [of_anti of_ident id] diff --git a/src/tac2quote.mli b/src/tac2quote.mli index 148e6818dd..403d333f38 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -80,6 +80,8 @@ val of_assertion : assertion -> raw_tacexpr val of_constr_matching : constr_matching -> raw_tacexpr +val of_goal_matching : goal_matching -> raw_tacexpr + (** {5 Generic arguments} *) val wit_pattern : (Constrexpr.constr_expr, Pattern.constr_pattern) Arg.tag diff --git a/theories/Notations.v b/theories/Notations.v index 91025ea964..8523c0f524 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -15,22 +15,20 @@ Ltac2 lazy_match0 t pats := let rec interp m := match m with | [] => Control.zero Match_failure | p :: m => - match p with - | Pattern.ConstrMatchPattern pat f => - Control.plus - (fun _ => - let bind := Pattern.matches_vect pat t in - fun _ => f bind - ) - (fun _ => interp m) - | Pattern.ConstrMatchContext pat f => - Control.plus - (fun _ => - let ((context, bind)) := Pattern.matches_subterm_vect pat t in - fun _ => f context bind - ) - (fun _ => interp m) - end + let next _ := interp m in + let ((knd, pat, f)) := p in + let p := match knd with + | Pattern.MatchPattern => + (fun _ => + let context := Pattern.empty_context () in + let bind := Pattern.matches_vect pat t in + fun _ => f context bind) + | Pattern.MatchContext => + (fun _ => + let ((context, bind)) := Pattern.matches_subterm_vect pat t in + fun _ => f context bind) + end in + Control.plus p next end in let ans := Control.once (fun () => interp pats) in ans (). @@ -42,22 +40,20 @@ Ltac2 multi_match0 t pats := let rec interp m := match m with | [] => Control.zero Match_failure | p :: m => - match p with - | Pattern.ConstrMatchPattern pat f => - Control.plus - (fun _ => - let bind := Pattern.matches_vect pat t in - f bind - ) - (fun _ => interp m) - | Pattern.ConstrMatchContext pat f => - Control.plus - (fun _ => - let ((context, bind)) := Pattern.matches_subterm_vect pat t in - f context bind - ) - (fun _ => interp m) - end + let next _ := interp m in + let ((knd, pat, f)) := p in + let p := match knd with + | Pattern.MatchPattern => + (fun _ => + let context := Pattern.empty_context () in + let bind := Pattern.matches_vect pat t in + f context bind) + | Pattern.MatchContext => + (fun _ => + let ((context, bind)) := Pattern.matches_subterm_vect pat t in + f context bind) + end in + Control.plus p next end in interp pats. diff --git a/theories/Pattern.v b/theories/Pattern.v index a672ad0fe7..fb450d9dfa 100644 --- a/theories/Pattern.v +++ b/theories/Pattern.v @@ -12,11 +12,15 @@ Ltac2 Type t := pattern. Ltac2 Type context. -Ltac2 Type 'a constr_match := [ -| ConstrMatchPattern (pattern, constr array -> 'a) -| ConstrMatchContext (pattern, context -> constr array -> 'a) +Ltac2 Type match_kind := [ +| MatchPattern +| MatchContext ]. +Ltac2 @ external empty_context : unit -> context := + "ltac2" "pattern_empty_context". +(** A trivial context only made of the hole. *) + Ltac2 @ external matches : t -> constr -> (ident * constr) list := "ltac2" "pattern_matches". (** If the term matches the pattern, returns the bound variables. If it doesn't, @@ -38,6 +42,25 @@ Ltac2 @ external matches_subterm_vect : t -> constr -> context * constr array := "ltac2" "pattern_matches_subterm_vect". (** Internal version of [matches_subterms] that does not return the identifiers. *) +Ltac2 @ external matches_goal : bool -> (match_kind * t) list -> (match_kind * t) -> + ident array * context array * constr array * context := + "ltac2" "pattern_matches_goal". +(** Given a list of patterns [hpats] for hypotheses and one pattern [cpat] for the + conclusion, [matches_goal rev hpats cpat] produces (a stream of) tuples of: + - An array of idents, whose size is the length of [hpats], corresponding to the + name of matched hypotheses. + - An array of contexts, whose size is the length of [hpats], corresponding to + the contexts matched for every hypothesis pattern. In case the match kind of + a hypothesis was [MatchPattern], the corresponding context is ensured to be empty. + - An array of terms, whose size is the total number of pattern variables without + duplicates. Terms are ordered by identifier order, e.g. ?a comes before ?b. + - A context corresponding to the conclusion, which is ensured to be empty if + the kind of [cpat] was [MatchPattern]. + This produces a backtracking stream of results containing all the possible + result combinations. The order of considered hypotheses is reversed if [rev] + is true. +*) + Ltac2 @ external instantiate : context -> constr -> constr := "ltac2" "pattern_instantiate". (** Fill the hole of a context with the given term. *) -- cgit v1.2.3 From 937b4a9ab459696fccd613e52601411c4f1dadef Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 27 Oct 2017 11:50:02 +0200 Subject: Fix relative meaning of Pattern vs. Context in match goal. --- src/tac2core.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index c4502f8eae..1cfb34c249 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -604,7 +604,7 @@ let () = define3 "pattern_matches_goal" bool (list (pair bool pattern)) (pair bo let env = Proofview.Goal.env gl in let sigma = Proofview.Goal.sigma gl in let concl = Proofview.Goal.concl gl in - let mk_pattern (b, pat) = if b then Tac2match.MatchContext pat else Tac2match.MatchPattern pat in + let mk_pattern (b, pat) = if b then Tac2match.MatchPattern pat else Tac2match.MatchContext pat in let r = (List.map mk_pattern hp, mk_pattern cp) in Tac2match.match_goal env sigma concl ~rev r >>= fun (hyps, ctx, subst) -> let of_ctxopt ctx = Value.of_constr (Option.default empty_context ctx) in -- cgit v1.2.3 From bd462a21974caca5928ae172a7740a1f96ae0ae4 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 27 Oct 2017 14:48:34 +0200 Subject: Fix goal_matching quotation. --- src/tac2quote.ml | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 9728c3af93..1275c939c5 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -419,10 +419,10 @@ let of_goal_matching (loc, gm) = in let map (loc, (pat, tac)) = let (pat, hyps, hctx, subst, cctx) = mk_gpat pat in - let tac = abstract_vars loc hyps tac in - let tac = abstract_vars loc hctx tac in + let tac = Loc.tag ?loc @@ CTacFun ([Loc.tag ?loc @@ CPatVar cctx, None], tac) in let tac = abstract_vars loc subst tac in - let tac = abstract_vars loc [cctx] tac in + let tac = abstract_vars loc hctx tac in + let tac = abstract_vars loc hyps tac in of_tuple ?loc [pat; tac] in of_list ?loc map gm -- cgit v1.2.3 From e0fd7c668bc284924c63a1f0a0e36fb4856c49e1 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 27 Oct 2017 16:04:00 +0200 Subject: Adding a command to evaluate Ltac2 expressions. --- doc/ltac2.md | 14 ++++++++++++++ src/g_ltac2.ml4 | 9 +++++++-- src/tac2entries.ml | 31 +++++++++++++++++++++++++++++++ src/tac2expr.mli | 2 ++ 4 files changed, 54 insertions(+), 2 deletions(-) diff --git a/doc/ltac2.md b/doc/ltac2.md index d7c8719a14..6cbe0988d0 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -727,6 +727,20 @@ foo 0 ↦ (fun x => x ()) (fun _ => 0) Note that abbreviations are not typechecked at all, and may result in typing errors after expansion. +# Evaluation + +Ltac2 features a toplevel loop that can be used to evaluate expressions. + +``` +VERNAC ::= +| "Ltac2" "Eval" TERM +``` + +This command evaluates the term in the current proof if there is one, or in the +global environment otherwise, and displays the resulting value to the user +together with its type. This function is pure in the sense that it does not +modify the state of the proof, and in particular all side-effects are discarded. + # Debug When the option `Ltac2 Backtrace` is set, toplevel failures will be printed with diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index be6c2de5a9..a979b1e9b8 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -88,6 +88,7 @@ let tac2def_typ = Gram.entry_create "tactic:tac2def_typ" let tac2def_ext = Gram.entry_create "tactic:tac2def_ext" let tac2def_syn = Gram.entry_create "tactic:tac2def_syn" let tac2def_mut = Gram.entry_create "tactic:tac2def_mut" +let tac2def_run = Gram.entry_create "tactic:tac2def_run" let tac2mode = Gram.entry_create "vernac:ltac2_command" (** FUCK YOU API *) @@ -109,7 +110,7 @@ let pattern_of_qualid ?loc id = GEXTEND Gram GLOBAL: tac2expr tac2type tac2def_val tac2def_typ tac2def_ext tac2def_syn - tac2def_mut; + tac2def_mut tac2def_run; tac2pat: [ "1" LEFTA [ id = Prim.qualid; pl = LIST1 tac2pat LEVEL "0" -> @@ -273,6 +274,9 @@ GEXTEND Gram tac2def_mut: [ [ "Set"; qid = Prim.qualid; ":="; e = tac2expr -> StrMut (qid, e) ] ] ; + tac2def_run: + [ [ "Eval"; e = tac2expr -> StrRun e ] ] + ; tac2typ_knd: [ [ t = tac2type -> CTydDef (Some t) | "["; ".."; "]" -> CTydOpn @@ -801,11 +805,12 @@ PRINTED BY pr_ltac2entry | [ tac2def_ext(e) ] -> [ e ] | [ tac2def_syn(e) ] -> [ e ] | [ tac2def_mut(e) ] -> [ e ] +| [ tac2def_run(e) ] -> [ e ] END let classify_ltac2 = function | StrSyn _ -> Vernacexpr.VtUnknown, Vernacexpr.VtNow -| StrMut _ | StrVal _ | StrPrm _ | StrTyp _ -> Vernac_classifier.classify_as_sideeff +| StrMut _ | StrVal _ | StrPrm _ | StrTyp _ | StrRun _ -> Vernac_classifier.classify_as_sideeff VERNAC COMMAND EXTEND VernacDeclareTactic2Definition | [ "Ltac2" ltac2_entry(e) ] => [ classify_ltac2 e ] -> [ diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 04bf21f656..b803278929 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -739,6 +739,36 @@ let register_redefinition ?(local = false) (loc, qid) e = } in Lib.add_anonymous_leaf (inTac2Redefinition def) +let perform_eval e = + let open Proofview.Notations in + let env = Global.env () in + let (e, ty) = Tac2intern.intern ~strict:false e in + let v = Tac2interp.interp Tac2interp.empty_environment e in + let selector, proof = + try + Proof_bullet.get_default_goal_selector (), + Proof_global.give_me_the_proof () + with Proof_global.NoCurrentProof -> + let sigma = Evd.from_env env in + Vernacexpr.SelectAll, Proof.start sigma [] + in + let v = match selector with + | Vernacexpr.SelectNth i -> Proofview.tclFOCUS i i v + | Vernacexpr.SelectList l -> Proofview.tclFOCUSLIST l v + | Vernacexpr.SelectId id -> Proofview.tclFOCUSID id v + | Vernacexpr.SelectAll -> v + in + (** HACK: the API doesn't allow to return a value *) + let ans = ref None in + let tac = (v >>= fun r -> ans := Some r; Proofview.tclUNIT ()) in + let (proof, _) = Proof.run_tactic (Global.env ()) tac proof in + let sigma = Proof.in_proof proof (fun sigma -> sigma) in + let ans = match !ans with None -> assert false | Some r -> r in + let name = int_name () in + Feedback.msg_notice (str "- : " ++ pr_glbtype name (snd ty) + ++ spc () ++ str "=" ++ spc () ++ + Tac2print.pr_valexpr env sigma ans (snd ty)) + (** Toplevel entries *) let register_struct ?local str = match str with @@ -747,6 +777,7 @@ let register_struct ?local str = match str with | StrPrm (id, t, ml) -> register_primitive ?local id t ml | StrSyn (tok, lev, e) -> register_notation ?local tok lev e | StrMut (qid, e) -> register_redefinition ?local qid e +| StrRun e -> perform_eval e (** Toplevel exception *) diff --git a/src/tac2expr.mli b/src/tac2expr.mli index e57b0ba3ef..89152dffe7 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -168,6 +168,8 @@ type strexpr = (** Syntactic extensions *) | StrMut of qualid located * raw_tacexpr (** Redefinition of mutable globals *) +| StrRun of raw_tacexpr + (** Toplevel evaluation of an expression *) (** {5 Dynamic semantics} *) -- cgit v1.2.3 From a7c83429db05866dfc9613fc4a488d62d31386fc Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 27 Oct 2017 18:42:02 +0200 Subject: Adding a notation for match goal. --- tests/matching.v | 25 +++++++++++++++++ theories/Notations.v | 71 +++++++++++++++------------------------------- theories/Pattern.v | 79 ++++++++++++++++++++++++++++++++++++++++++++++++++++ 3 files changed, 127 insertions(+), 48 deletions(-) diff --git a/tests/matching.v b/tests/matching.v index f43e0121ef..6bc5706da7 100644 --- a/tests/matching.v +++ b/tests/matching.v @@ -1,5 +1,12 @@ Require Import Ltac2.Ltac2 Ltac2.Notations. +Ltac2 Type exn ::= [ Nope ]. + +Ltac2 check_id id id' := match Ident.equal id id' with +| true => () +| false => Control.throw Nope +end. + Goal True -> False. Proof. Fail @@ -25,3 +32,21 @@ let f c := in match! '(nat -> bool) with context [?a] => f a end. Abort. + +Goal forall (i j : unit) (x y : nat) (b : bool), True. +Proof. +Fail match! goal with +| [ h : ?t, h' : ?t |- _ ] => () +end. +intros i j x y b. +match! goal with +| [ h : ?t, h' : ?t |- _ ] => + check_id h @x; + check_id h' @y +end. +match! reverse goal with +| [ h : ?t, h' : ?t |- _ ] => + check_id h @j; + check_id h' @i +end. +Abort. diff --git a/theories/Notations.v b/theories/Notations.v index 8523c0f524..48f3ca0587 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -11,59 +11,34 @@ Require Ltac2.Control Ltac2.Pattern Ltac2.Array Ltac2.Int Ltac2.Std. (** Constr matching *) -Ltac2 lazy_match0 t pats := - let rec interp m := match m with - | [] => Control.zero Match_failure - | p :: m => - let next _ := interp m in - let ((knd, pat, f)) := p in - let p := match knd with - | Pattern.MatchPattern => - (fun _ => - let context := Pattern.empty_context () in - let bind := Pattern.matches_vect pat t in - fun _ => f context bind) - | Pattern.MatchContext => - (fun _ => - let ((context, bind)) := Pattern.matches_subterm_vect pat t in - fun _ => f context bind) - end in - Control.plus p next - end in - let ans := Control.once (fun () => interp pats) in - ans (). - Ltac2 Notation "lazy_match!" t(tactic(6)) "with" m(constr_matching) "end" := - lazy_match0 t m. - -Ltac2 multi_match0 t pats := - let rec interp m := match m with - | [] => Control.zero Match_failure - | p :: m => - let next _ := interp m in - let ((knd, pat, f)) := p in - let p := match knd with - | Pattern.MatchPattern => - (fun _ => - let context := Pattern.empty_context () in - let bind := Pattern.matches_vect pat t in - f context bind) - | Pattern.MatchContext => - (fun _ => - let ((context, bind)) := Pattern.matches_subterm_vect pat t in - f context bind) - end in - Control.plus p next - end in - interp pats. + Pattern.lazy_match0 t m. Ltac2 Notation "multi_match!" t(tactic(6)) "with" m(constr_matching) "end" := - multi_match0 t m. - -Ltac2 one_match0 t m := Control.once (fun _ => multi_match0 t m). + Pattern.multi_match0 t m. Ltac2 Notation "match!" t(tactic(6)) "with" m(constr_matching) "end" := - one_match0 t m. + Pattern.one_match0 t m. + +(** Goal matching *) + +Ltac2 Notation "lazy_match!" "goal" "with" m(goal_matching) "end" := + Pattern.lazy_goal_match0 false m. + +Ltac2 Notation "multi_match!" "goal" "with" m(goal_matching) "end" := + Pattern.multi_goal_match0 false m. + +Ltac2 Notation "match!" "goal" "with" m(goal_matching) "end" := + Pattern.one_goal_match0 false m. + +Ltac2 Notation "lazy_match!" "reverse" "goal" "with" m(goal_matching) "end" := + Pattern.lazy_goal_match0 true m. + +Ltac2 Notation "multi_match!" "reverse" "goal" "with" m(goal_matching) "end" := + Pattern.multi_goal_match0 true m. + +Ltac2 Notation "match!" "reverse" "goal" "with" m(goal_matching) "end" := + Pattern.one_goal_match0 true m. (** Tacticals *) diff --git a/theories/Pattern.v b/theories/Pattern.v index fb450d9dfa..2c93918aee 100644 --- a/theories/Pattern.v +++ b/theories/Pattern.v @@ -7,6 +7,7 @@ (************************************************************************) Require Import Ltac2.Init. +Require Ltac2.Control. Ltac2 Type t := pattern. @@ -64,3 +65,81 @@ Ltac2 @ external matches_goal : bool -> (match_kind * t) list -> (match_kind * t Ltac2 @ external instantiate : context -> constr -> constr := "ltac2" "pattern_instantiate". (** Fill the hole of a context with the given term. *) + +(** Implementation of Ltac matching over terms and goals *) + +Ltac2 lazy_match0 t pats := + let rec interp m := match m with + | [] => Control.zero Match_failure + | p :: m => + let next _ := interp m in + let ((knd, pat, f)) := p in + let p := match knd with + | MatchPattern => + (fun _ => + let context := empty_context () in + let bind := matches_vect pat t in + fun _ => f context bind) + | MatchContext => + (fun _ => + let ((context, bind)) := matches_subterm_vect pat t in + fun _ => f context bind) + end in + Control.plus p next + end in + Control.once (fun () => interp pats) (). + +Ltac2 multi_match0 t pats := + let rec interp m := match m with + | [] => Control.zero Match_failure + | p :: m => + let next _ := interp m in + let ((knd, pat, f)) := p in + let p := match knd with + | MatchPattern => + (fun _ => + let context := empty_context () in + let bind := matches_vect pat t in + f context bind) + | MatchContext => + (fun _ => + let ((context, bind)) := matches_subterm_vect pat t in + f context bind) + end in + Control.plus p next + end in + interp pats. + +Ltac2 one_match0 t m := Control.once (fun _ => multi_match0 t m). + +Ltac2 lazy_goal_match0 rev pats := + let rec interp m := match m with + | [] => Control.zero Match_failure + | p :: m => + let next _ := interp m in + let ((pat, f)) := p in + let ((phyps, pconcl)) := pat in + let cur _ := + let ((hids, hctx, subst, cctx)) := matches_goal rev phyps pconcl in + fun _ => f hids hctx subst cctx + in + Control.plus cur next + end in + interp pats (). + +Ltac2 multi_goal_match0 rev pats := + let rec interp m := match m with + | [] => Control.zero Match_failure + | p :: m => + let next _ := interp m in + let ((pat, f)) := p in + let ((phyps, pconcl)) := pat in + let cur _ := + let ((hids, hctx, subst, cctx)) := matches_goal rev phyps pconcl in + f hids hctx subst cctx + in + Control.plus cur next + end in + interp pats. + +Ltac2 one_goal_match0 rev pats := Control.once (fun _ => multi_goal_match0 rev pats). -- cgit v1.2.3 From 216c5f25cf41d68871149f21f83518ec0a4f1cc9 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 27 Oct 2017 19:26:31 +0200 Subject: Adding an OPAM package. --- Makefile | 16 +++++++--------- opam/descr | 1 + opam/opam | 17 +++++++++++++++++ 3 files changed, 25 insertions(+), 9 deletions(-) create mode 100644 opam/descr create mode 100644 opam/opam diff --git a/Makefile b/Makefile index d555fea236..e0e197650d 100644 --- a/Makefile +++ b/Makefile @@ -1,16 +1,14 @@ -all: Makefile.coq - $(MAKE) -f Makefile.coq +ifeq "$(COQBIN)" "" + COQBIN=$(dir $(shell which coqtop))/ +endif -install: all - $(MAKE) -f Makefile.coq install - -clean: Makefile.coq - $(MAKE) -f Makefile.coq clean - rm -f Makefile.coq +%: Makefile.coq Makefile.coq: _CoqProject - $(COQBIN)/coq_makefile -f _CoqProject -o Makefile.coq + $(COQBIN)coq_makefile -f _CoqProject -o Makefile.coq tests: all @$(MAKE) -C tests -s clean @$(MAKE) -C tests -s all + +-include Makefile.coq diff --git a/opam/descr b/opam/descr new file mode 100644 index 0000000000..82463c4f45 --- /dev/null +++ b/opam/descr @@ -0,0 +1 @@ +A tactic language for Coq. diff --git a/opam/opam b/opam/opam new file mode 100644 index 0000000000..e461b97942 --- /dev/null +++ b/opam/opam @@ -0,0 +1,17 @@ +opam-version: "1.2" +name: "coq-ltac2" +version: "0.1" +maintainer: "Pierre-Marie Pédrot " +author: "Pierre-Marie Pédrot " +license: "LGPL 2.1" +homepage: "https://github.com/ppedrot/ltac2" +dev-repo: "https://github.com/ppedrot/ltac2.git" +bug-reports: "https://github.com/ppedrot/ltac2/issues" +build: [ + [make "COQBIN=\"\"" "-j%{jobs}%"] +] +install: [make "install"] +remove: [make "uninstall"] +depends: [ + "coq" { = "dev" } +] -- cgit v1.2.3 From e6efa6c12dd4701dc7fbdd31580bad0ad676e30d Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 27 Oct 2017 19:45:29 +0200 Subject: Better printers for toplevel values. --- src/tac2print.ml | 53 ++++++++++++++++++++++++++++++++++++++++++----------- 1 file changed, 42 insertions(+), 11 deletions(-) diff --git a/src/tac2print.ml b/src/tac2print.ml index 8e4947e332..8f61686988 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -320,7 +320,7 @@ let pr_glbexpr_gen lvl c = let kn = change_kn_label tpe id in pr_projection kn ++ spc () ++ str ":=" ++ spc () ++ pr_glbexpr E1 arg in - let args = prlist_with_sep (fun () -> str ";" ++ spc ()) pr_arg args in + let args = prlist_with_sep pr_semicolon pr_arg args in hv 0 (str "{" ++ spc () ++ args ++ spc () ++ str "}") | (GTydDef _ | GTydOpn) -> assert false in @@ -374,7 +374,9 @@ let rec pr_valexpr env sigma v t = match kind t with | Some pr -> pr.val_printer env sigma v params | None -> let n, repr = Tac2env.interp_type kn in - match repr with + if KerName.equal kn t_list then + pr_val_list env sigma (to_list (fun v -> repr_to valexpr v) v) (List.hd params) + else match repr with | GTydDef None -> str "" | GTydDef (Some _) -> (** Shouldn't happen thanks to kind *) @@ -388,7 +390,7 @@ let rec pr_valexpr env sigma v t = match kind t with let knc = change_kn_label kn id in let args = pr_constrargs env sigma params args tpe in hv 2 (pr_constructor knc ++ spc () ++ str "(" ++ args ++ str ")") - | GTydRec rcd -> str "{}" + | GTydRec rcd -> str "{ TODO }" | GTydOpn -> begin match Tac2ffi.to_open v with | (knc, [||]) -> pr_constructor knc @@ -399,6 +401,7 @@ let rec pr_valexpr env sigma v t = match kind t with end end | GTypArrow _ -> str "" +| GTypRef (Tuple 0, []) -> str "()" | GTypRef (Tuple _, tl) -> let blk = Array.to_list (snd (to_block v)) in if List.length blk == List.length tl then @@ -414,31 +417,59 @@ and pr_constrargs env sigma params args tpe = let args = List.combine args tpe in prlist_with_sep pr_comma (fun (v, t) -> pr_valexpr env sigma v t) args +and pr_val_list env sigma args tpe = + let pr v = pr_valexpr env sigma v tpe in + str "[" ++ prlist_with_sep pr_semicolon pr args ++ str "]" + let register_init n f = let kn = KerName.make2 Tac2env.coq_prefix (Label.make n) in - register_val_printer kn { val_printer = fun _ _ v _ -> f v } + register_val_printer kn { val_printer = fun env sigma v _ -> f env sigma v } -let () = register_init "int" begin fun n -> +let () = register_init "int" begin fun _ _ n -> let n = to_int n in Pp.int n end -let () = register_init "string" begin fun s -> +let () = register_init "string" begin fun _ _ s -> let s = to_string s in - Pp.quote (Pp.str (Bytes.to_string s)) + Pp.quote (str (Bytes.to_string s)) end -let () = register_init "ident" begin fun id -> +let () = register_init "ident" begin fun _ _ id -> let id = to_ident id in - Pp.str "@" ++ Id.print id + str "@" ++ Id.print id +end + +let () = register_init "constr" begin fun env sigma c -> + let c = to_constr c in + let c = try Printer.pr_leconstr_env env sigma c with _ -> str "..." in + str "constr:(" ++ c ++ str ")" +end + +let () = register_init "pattern" begin fun env sigma c -> + let c = to_pattern c in + let c = try Printer.pr_lconstr_pattern_env env sigma c with _ -> str "..." in + str "pattern:(" ++ c ++ str ")" end -let () = register_init "message" begin fun pp -> +let () = register_init "message" begin fun _ _ pp -> str "message:(" ++ to_pp pp ++ str ")" end -let () = register_init "err" begin fun e -> +let () = register_init "err" begin fun _ _ e -> let e = to_ext val_exn e in let (e, _) = ExplainErr.process_vernac_interp_error ~allow_uncaught:true e in str "err:(" ++ CErrors.print_no_report e ++ str ")" end + +let () = + let kn = KerName.make2 Tac2env.coq_prefix (Label.make "array") in + let val_printer env sigma v arg = match arg with + | [arg] -> + let (_, v) = to_block v in + str "[|" ++ spc () ++ + prvect_with_sep pr_semicolon (fun a -> pr_valexpr env sigma a arg) v ++ + spc () ++ str "|]" + | _ -> assert false + in + register_val_printer kn { val_printer } -- cgit v1.2.3 From 71208e3eee6745ed8849bd03f66db638d9897516 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 27 Oct 2017 21:07:50 +0200 Subject: Adding documentation --- doc/ltac2.md | 91 ++++++++++++++++++++++++++++++++++++++++++++++++++++++------ 1 file changed, 83 insertions(+), 8 deletions(-) diff --git a/doc/ltac2.md b/doc/ltac2.md index 6cbe0988d0..cd0d8f4325 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -268,7 +268,7 @@ not to compute the argument, and `foo` would have e.g. type `(unit -> unit) -> unit`. ``` -foo (fun () -> let x := 0 in bar) +foo (fun () => let x := 0 in bar) ``` ## Typing @@ -559,7 +559,7 @@ for it. - `&x` as an Ltac2 expression expands to `hyp @x`. - `&x` as a Coq constr expression expands to - `ltac2:(refine (fun () => hyp @x))`. + `ltac2:(Control.refine (fun () => hyp @x))`. #### Dynamic semantics @@ -587,14 +587,86 @@ Many standard tactics perform type-checking of their argument before going further. It is your duty to ensure that terms are well-typed when calling such tactics. Failure to do so will result in non-recoverable exceptions. -## Patterns +## Trivial Term Antiquotations -Terms can be used in pattern position just as any Ltac constructor. The accepted -syntax is a subset of the constr syntax in Ltac term position. It does not -allow antiquotations. +It is possible to refer to a variable of type `constr` in the Ltac2 environment +through a specific syntax consistent with the antiquotations presented in +the notation section. -Patterns quotations are typically used with the matching functions provided -in the `Pattern` module. +``` +COQCONSTR ::= +| "$" LIDENT +``` + +In a Coq term, writing `$x` is semantically equivalent to +`ltac2:(Control.refine (fun () => x))`, up to re-typechecking. It allows to +insert in a concise way an Ltac2 variable of type `constr` into a Coq term. + +## Match over terms + +Ltac2 features a construction similar to Ltac1 `match` over terms, although +in a less hard-wired way. + +``` +TERM ::= +| "match!" TERM "with" CONSTR-MATCHING* "end" +| "lazy_match!" TERM "with" CONSTR-MATCHING* "end" +| "multi_match!" TERM "with" CONSTR-MATCHING*"end" + +CONSTR-MATCHING := +| "|" CONSTR-PATTERN "=>" TERM + +CONSTR-PATTERN := +| CONSTR +| "context" LIDENT? "[" CONSTR "]" +``` + +This construction is not primitive and is desugared at parsing time into +calls to term matching functions from the `Pattern` module. Internally, it is +implemented thanks to a specific scope accepting the `CONSTR-MATCHING` syntax. + +Variables from the `CONSTR-PATTERN` are statically bound in the body of the branch, to +values of type `constr` for the variables from the `CONSTR` pattern and to a +value of type `Pattern.context` for the variable `LIDENT`. + +Note that contrarily to Ltac, only lowercase identifiers are valid as Ltac2 +bindings, so that there will be a syntax error if one of the bound variables +starts with an uppercase character. + +The semantics of this construction is otherwise the same as the corresponding +one from Ltac1, except that it requires the goal to be focussed. + +## Match over goals + +Similarly, there is a way to match over goals in an elegant way, which is +just a notation desugared at parsing time. + +``` +TERM ::= +| "match!" MATCH-ORDER? "goal" "with" GOAL-MATCHING* "end" +| "lazy_match!" MATCH-ORDER? "goal" "with" GOAL-MATCHING* "end" +| "multi_match!" MATCH-ORDER? "goal" "with" GOAL-MATCHING*"end" + +GOAL-MATCHING := +| "|" "[" HYP-MATCHING* "|-" CONSTR-PATTERN "]" "=>" TERM + +HYP-MATCHING := +| LIDENT ":" CONSTR-PATTERN + +MATCH-ORDER := +| "reverse" +``` + +Variables from `HYP-MATCHING` and `CONSTR-PATTERN` are bound in the body of the +branch. Their types are: +- `constr` for pattern variables appearing in a `CONSTR` +- `Pattern.context` for variables binding a context +- `ident` for variables binding a hypothesis name. + +The same identifier caveat as in the case of matching over constr applies, and +this features has the same semantics as in Ltac1. In particular, a `reverse` +flag can be specified to match hypotheses from the more recently introduced to +the least recently introduced one. # Notations @@ -656,6 +728,9 @@ The following scopes are built-in. of returning a useless unit value. It is forbidden for the various subscopes to refer to the global entry using self of next. +A few other specific scopes exist to handle Ltac1-like syntax, but their use is +discouraged and they are thus not documented. + For now there is no way to declare new scopes from Ltac2 side, but this is planned. -- cgit v1.2.3 From f18502f32fb25b29cafe26340edbbcedd463c646 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 30 Oct 2017 11:40:49 +0100 Subject: Fix compilation after merge of Ltac_pretype interface. --- src/tac2core.ml | 2 +- src/tac2interp.mli | 4 ++-- src/tac2match.ml | 2 +- src/tac2match.mli | 2 +- src/tac2tactics.ml | 2 +- 5 files changed, 6 insertions(+), 6 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 1cfb34c249..9ef88a7d56 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -828,7 +828,7 @@ let open_constr_no_classes_flags () = let to_lvar ist = let open Glob_ops in let lfun = Tac2interp.set_env ist Id.Map.empty in - { empty_lvar with Glob_term.ltac_genargs = lfun } + { empty_lvar with Ltac_pretype.ltac_genargs = lfun } let gtypref kn = GTypRef (Other kn, []) diff --git a/src/tac2interp.mli b/src/tac2interp.mli index 211ac95196..21fdcd03af 100644 --- a/src/tac2interp.mli +++ b/src/tac2interp.mli @@ -20,8 +20,8 @@ val interp : environment -> glb_tacexpr -> valexpr Proofview.tactic (** {5 Cross-boundary encodings} *) -val get_env : Glob_term.unbound_ltac_var_map -> environment -val set_env : environment -> Glob_term.unbound_ltac_var_map -> Glob_term.unbound_ltac_var_map +val get_env : Ltac_pretype.unbound_ltac_var_map -> environment +val set_env : environment -> Ltac_pretype.unbound_ltac_var_map -> Ltac_pretype.unbound_ltac_var_map (** {5 Exceptions} *) diff --git a/src/tac2match.ml b/src/tac2match.ml index 7a22e91b6b..fef5647725 100644 --- a/src/tac2match.ml +++ b/src/tac2match.ml @@ -14,7 +14,7 @@ module NamedDecl = Context.Named.Declaration type context = EConstr.t type result = { - subst : Pattern.patvar_map ; + subst : Ltac_pretype.patvar_map ; } type match_pattern = diff --git a/src/tac2match.mli b/src/tac2match.mli index cf64542f40..7cfa1ed25f 100644 --- a/src/tac2match.mli +++ b/src/tac2match.mli @@ -30,4 +30,4 @@ val match_goal: match_rule -> ((Id.t * context option) list * (** List of hypotheses matching: name + context *) context option * (** Context for conclusion *) - Pattern.patvar_map (** Pattern variable substitution *)) Proofview.tactic + Ltac_pretype.patvar_map (** Pattern variable substitution *)) Proofview.tactic diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index b496f5046f..c68fdf9a7a 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -343,7 +343,7 @@ let discriminate ev arg = let injection ev ipat arg = let arg = Option.map (fun arg -> None, arg) arg in let ipat = Option.map mk_intro_patterns ipat in - let tac ev arg = Equality.injClause ipat ev arg in + let tac ev arg = Equality.injClause None ipat ev arg in on_destruction_arg tac ev arg let autorewrite ~all by ids cl = -- cgit v1.2.3 From a997ee7d78d90740b15b58502a1dc5e587b43ee3 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 30 Oct 2017 15:53:55 +0100 Subject: Introducing the change tactic. --- src/g_ltac2.ml4 | 12 ++++++- src/tac2core.ml | 1 + src/tac2entries.ml | 1 + src/tac2entries.mli | 1 + src/tac2qexpr.mli | 6 ++++ src/tac2quote.ml | 98 ++++++++++++++++++++++++++++++---------------------- src/tac2quote.mli | 2 ++ src/tac2stdlib.ml | 7 ++++ src/tac2tactics.ml | 12 +++++++ src/tac2tactics.mli | 2 ++ tests/example2.v | 6 ++++ theories/Notations.v | 6 ++++ theories/Std.v | 2 ++ 13 files changed, 113 insertions(+), 43 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index a979b1e9b8..557e2bcb9a 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -378,7 +378,7 @@ let loc_of_ne_list l = Loc.merge_opt (fst (List.hd l)) (fst (List.last l)) GEXTEND Gram GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause - q_rewriting q_clause q_dispatch q_occurrences q_strategy_flag + q_conversion q_rewriting q_clause q_dispatch q_occurrences q_strategy_flag q_destruction_arg q_reference q_with_bindings q_constr_matching q_goal_matching q_hintdb q_move_location q_pose q_assert; anti: @@ -585,6 +585,16 @@ GEXTEND Gram q_induction_clause: [ [ cl = induction_clause -> cl ] ] ; + conversion: + [ [ c = Constr.constr -> + Loc.tag ~loc:!@loc @@ QConvert c + | c1 = Constr.constr; "with"; c2 = Constr.constr -> + Loc.tag ~loc:!@loc @@ QConvertWith (c1, c2) + ] ] + ; + q_conversion: + [ [ c = conversion -> c ] ] + ; orient: [ [ "->" -> Loc.tag ~loc:!@loc (Some true) | "<-" -> Loc.tag ~loc:!@loc (Some false) diff --git a/src/tac2core.ml b/src/tac2core.ml index 9ef88a7d56..ce53b781f5 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -1160,6 +1160,7 @@ let () = add_expr_scope "intropattern" q_intropattern Tac2quote.of_intro_pattern let () = add_expr_scope "intropatterns" q_intropatterns Tac2quote.of_intro_patterns let () = add_expr_scope "destruction_arg" q_destruction_arg Tac2quote.of_destruction_arg let () = add_expr_scope "induction_clause" q_induction_clause Tac2quote.of_induction_clause +let () = add_expr_scope "conversion" q_conversion Tac2quote.of_conversion let () = add_expr_scope "rewriting" q_rewriting Tac2quote.of_rewriting let () = add_expr_scope "clause" q_clause Tac2quote.of_clause let () = add_expr_scope "hintdb" q_hintdb Tac2quote.of_hintdb diff --git a/src/tac2entries.ml b/src/tac2entries.ml index b803278929..3ac3d14ef3 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -30,6 +30,7 @@ let q_intropattern = Pcoq.Gram.entry_create "tactic:q_intropattern" let q_intropatterns = Pcoq.Gram.entry_create "tactic:q_intropatterns" let q_destruction_arg = Pcoq.Gram.entry_create "tactic:q_destruction_arg" let q_induction_clause = Pcoq.Gram.entry_create "tactic:q_induction_clause" +let q_conversion = Pcoq.Gram.entry_create "tactic:q_conversion" let q_rewriting = Pcoq.Gram.entry_create "tactic:q_rewriting" let q_clause = Pcoq.Gram.entry_create "tactic:q_clause" let q_dispatch = Pcoq.Gram.entry_create "tactic:q_dispatch" diff --git a/src/tac2entries.mli b/src/tac2entries.mli index b2a2dd4846..a92e149a85 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -70,6 +70,7 @@ val q_intropattern : intro_pattern Pcoq.Gram.entry val q_intropatterns : intro_pattern list located Pcoq.Gram.entry val q_destruction_arg : destruction_arg Pcoq.Gram.entry val q_induction_clause : induction_clause Pcoq.Gram.entry +val q_conversion : conversion Pcoq.Gram.entry val q_rewriting : rewriting Pcoq.Gram.entry val q_clause : clause Pcoq.Gram.entry val q_dispatch : dispatch Pcoq.Gram.entry diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index 229cece7c4..ad52884ca6 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -84,6 +84,12 @@ type induction_clause_r = { type induction_clause = induction_clause_r located +type conversion_r = +| QConvert of Constrexpr.constr_expr +| QConvertWith of Constrexpr.constr_expr * Constrexpr.constr_expr + +type conversion = conversion_r located + type multi_r = | QPrecisely of int located | QUpTo of int located diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 1275c939c5..399c1199bd 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -208,6 +208,62 @@ let of_induction_clause (loc, cl) = std_proj "indcl_in", in_; ]) +let check_pattern_id ?loc id = + if Tac2env.is_constructor (Libnames.qualid_of_ident id) then + CErrors.user_err ?loc (str "Invalid pattern binding name " ++ Id.print id) + +let pattern_vars pat = + let rec aux () accu pat = match pat.CAst.v with + | Constrexpr.CPatVar id + | Constrexpr.CEvar (id, []) -> + let () = check_pattern_id ?loc:pat.CAst.loc id in + Id.Set.add id accu + | _ -> + Topconstr.fold_constr_expr_with_binders (fun _ () -> ()) aux () accu pat + in + aux () Id.Set.empty pat + +let abstract_vars loc vars tac = + let get_name = function Name id -> Some id | Anonymous -> None in + let def = try Some (List.find_map get_name vars) with Not_found -> None in + let na, tac = match def with + | None -> (Anonymous, tac) + | Some id0 -> + (** Trick: in order not to shadow a variable nor to choose an arbitrary + name, we reuse one which is going to be shadowed by the matched + variables anyways. *) + let build_bindings (n, accu) na = match na with + | Anonymous -> (n + 1, accu) + | Name _ -> + let get = global_ref ?loc (kername array_prefix "get") in + let args = [of_variable (loc, id0); of_int (loc, n)] in + let e = Loc.tag ?loc @@ CTacApp (get, args) in + let accu = (Loc.tag ?loc @@ CPatVar na, None, e) :: accu in + (n + 1, accu) + in + let (_, bnd) = List.fold_left build_bindings (0, []) vars in + let tac = Loc.tag ?loc @@ CTacLet (false, bnd, tac) in + (Name id0, tac) + in + Loc.tag ?loc @@ CTacFun ([Loc.tag ?loc @@ CPatVar na, None], tac) + +let of_pattern p = + inj_wit ?loc:p.CAst.loc wit_pattern p + +let of_conversion (loc, c) = match c with +| QConvert c -> + let pat = of_option ?loc of_pattern None in + let c = Loc.tag ?loc @@ CTacFun ([Loc.tag ?loc @@ CPatVar Anonymous, None], of_constr c) in + of_tuple ?loc [pat; c] +| QConvertWith (pat, c) -> + let vars = pattern_vars pat in + let pat = of_option ?loc of_pattern (Some pat) in + let c = of_constr c in + (** Order is critical here *) + let vars = List.map (fun id -> Name id) (Id.Set.elements vars) in + let c = abstract_vars loc vars c in + of_tuple [pat; c] + let of_repeat (loc, r) = match r with | QPrecisely n -> std_constructor ?loc "Precisely" [of_int n] | QUpTo n -> std_constructor ?loc "UpTo" [of_int n] @@ -307,45 +363,6 @@ let of_hintdb (loc, hdb) = match hdb with | QHintAll -> of_option ?loc (fun l -> of_list (fun id -> of_anti of_ident id) l) None | QHintDbs ids -> of_option ?loc (fun l -> of_list (fun id -> of_anti of_ident id) l) (Some ids) -let check_pattern_id ?loc id = - if Tac2env.is_constructor (Libnames.qualid_of_ident id) then - CErrors.user_err ?loc (str "Invalid pattern binding name " ++ Id.print id) - -let pattern_vars pat = - let rec aux () accu pat = match pat.CAst.v with - | Constrexpr.CPatVar id - | Constrexpr.CEvar (id, []) -> - let () = check_pattern_id ?loc:pat.CAst.loc id in - Id.Set.add id accu - | _ -> - Topconstr.fold_constr_expr_with_binders (fun _ () -> ()) aux () accu pat - in - aux () Id.Set.empty pat - -let abstract_vars loc vars tac = - let get_name = function Name id -> Some id | Anonymous -> None in - let def = try Some (List.find_map get_name vars) with Not_found -> None in - let na, tac = match def with - | None -> (Anonymous, tac) - | Some id0 -> - (** Trick: in order not to shadow a variable nor to choose an arbitrary - name, we reuse one which is going to be shadowed by the matched - variables anyways. *) - let build_bindings (n, accu) na = match na with - | Anonymous -> (n + 1, accu) - | Name _ -> - let get = global_ref ?loc (kername array_prefix "get") in - let args = [of_variable (loc, id0); of_int (loc, n)] in - let e = Loc.tag ?loc @@ CTacApp (get, args) in - let accu = (Loc.tag ?loc @@ CPatVar na, None, e) :: accu in - (n + 1, accu) - in - let (_, bnd) = List.fold_left build_bindings (0, []) vars in - let tac = Loc.tag ?loc @@ CTacLet (false, bnd, tac) in - (Name id0, tac) - in - Loc.tag ?loc @@ CTacFun ([Loc.tag ?loc @@ CPatVar na, None], tac) - let extract_name ?loc oid = match oid with | None -> Anonymous | Some id -> @@ -378,9 +395,6 @@ let of_constr_matching (loc, m) = in of_list ?loc map m -let of_pattern p = - inj_wit ?loc:p.CAst.loc wit_pattern p - (** From the patterns and the body of the branch, generate: - a goal pattern: (constr_match list * constr_match) - a branch function (ident array -> context array -> constr array -> context -> 'a) diff --git a/src/tac2quote.mli b/src/tac2quote.mli index 403d333f38..3f6c9a55e5 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -51,6 +51,8 @@ val of_destruction_arg : destruction_arg -> raw_tacexpr val of_induction_clause : induction_clause -> raw_tacexpr +val of_conversion : conversion -> raw_tacexpr + val of_rewriting : rewriting -> raw_tacexpr val of_occurrences : occurrences -> raw_tacexpr diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 99fa0370e1..60b6e70d58 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -467,6 +467,13 @@ let () = define_red2 "eval_native" begin fun where c -> Tac2tactics.eval_native where c end +let () = define_prim3 "tac_change" begin fun pat c cl -> + let pat = Value.to_option (fun p -> Value.to_pattern p) pat in + let c = Value.to_fun1 (array constr) constr c in + let cl = to_clause cl in + Tac2tactics.change pat c cl +end + let () = define_prim4 "tac_rewrite" begin fun ev rw cl by -> let ev = Value.to_bool ev in let rw = Value.to_list to_rewriting rw in diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index c68fdf9a7a..5a5b259ee7 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -150,6 +150,18 @@ let split_with_bindings ev bnd = let bnd = mk_bindings bnd in Tactics.split_with_bindings ev [bnd] +let change pat c cl = + let open Tac2ffi in + Proofview.Goal.enter begin fun gl -> + let env = Proofview.Goal.env gl in + let c subst sigma = + let subst = Array.map_of_list snd (Id.Map.bindings subst) in + delayed_of_tactic (Tac2ffi.app_fun1 c (array constr) constr subst) env sigma + in + let cl = mk_clause cl in + Tactics.change pat c cl + end + let rewrite ev rw cl by = let map_rw (orient, repeat, c) = let c = c >>= fun c -> return (mk_with_bindings c) in diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli index 3d64e7ec8c..7a4624ba2c 100644 --- a/src/tac2tactics.mli +++ b/src/tac2tactics.mli @@ -38,6 +38,8 @@ val left_with_bindings : evars_flag -> bindings -> unit tactic val right_with_bindings : evars_flag -> bindings -> unit tactic val split_with_bindings : evars_flag -> bindings -> unit tactic +val change : Pattern.constr_pattern option -> (constr array, constr) Tac2ffi.fun1 -> clause -> unit tactic + val rewrite : evars_flag -> rewriting list -> clause -> unit tactic option -> unit tactic diff --git a/tests/example2.v b/tests/example2.v index 20819606db..46e4e43ed0 100644 --- a/tests/example2.v +++ b/tests/example2.v @@ -260,3 +260,9 @@ Proof. assert (H : 0 + 0 = 0) by reflexivity. intros x; exact x. Qed. + +Goal 1 + 1 = 2. +Proof. +change (?a + 1 = 2) with (2 = $a + 1). +reflexivity. +Qed. diff --git a/theories/Notations.v b/theories/Notations.v index 48f3ca0587..77f3e235f3 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -375,6 +375,12 @@ Ltac2 Notation "native_compute" pl(opt(seq(pattern, occurrences))) cl(opt(clause Std.native pl (default_on_concl cl). Ltac2 Notation native_compute := native_compute. +Ltac2 change0 p cl := + let ((pat, c)) := p in + Std.change pat c (default_on_concl cl). + +Ltac2 Notation "change" c(conversion) cl(opt(clause)) := change0 c cl. + Ltac2 rewrite0 ev rw cl tac := let cl := default_on_concl cl in Std.rewrite ev rw cl tac. diff --git a/theories/Std.v b/theories/Std.v index 7831baf046..389299f266 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -175,6 +175,8 @@ Ltac2 @ external eval_pattern : (constr * occurrences) list -> constr -> constr Ltac2 @ external eval_vm : (pattern * occurrences) option -> constr -> constr := "ltac2" "eval_vm". Ltac2 @ external eval_native : (pattern * occurrences) option -> constr -> constr := "ltac2" "eval_native". +Ltac2 @ external change : pattern option -> (constr array -> constr) -> clause -> unit := "ltac2" "tac_change". + Ltac2 @ external rewrite : evar_flag -> rewriting list -> clause -> (unit -> unit) option -> unit := "ltac2" "tac_rewrite". Ltac2 @ external reflexivity : unit -> unit := "ltac2" "tac_reflexivity". -- cgit v1.2.3 From 2890c90d0d5900db4c47398e1f809f3c759e07e0 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 30 Oct 2017 17:55:02 +0100 Subject: Fix the semantics of introducing with empty intro patterns. --- src/tac2tactics.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index 5a5b259ee7..f7ad057e86 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -88,7 +88,7 @@ let mk_clause cl = { let intros_patterns ev ipat = let ipat = mk_intro_patterns ipat in - Tactics.intro_patterns ev ipat + Tactics.intros_patterns ev ipat let apply adv ev cb cl = let map c = -- cgit v1.2.3 From 2f95db622428cf58968e8d13f2230fe51c31e97f Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 30 Oct 2017 18:11:26 +0100 Subject: Put notations in level 5 by default. --- src/tac2entries.ml | 1 + 1 file changed, 1 insertion(+) diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 3ac3d14ef3..cefb4b13b8 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -674,6 +674,7 @@ let register_notation ?(local = false) tkn lev body = match tkn, lev with let ids = List.fold_left fold Id.Set.empty entries in (** Globalize so that names are absolute *) let body = Tac2intern.globalize ids body in + let lev = match lev with Some _ -> lev | None -> Some 5 in let ext = { synext_tok = tkn; synext_exp = body; -- cgit v1.2.3 From de7483beb78e5bd81dc6449ba201fb9dfc490ba8 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 30 Oct 2017 18:16:18 +0100 Subject: Add a macro for the now tactic --- theories/Notations.v | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/theories/Notations.v b/theories/Notations.v index 77f3e235f3..dc09812254 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -555,3 +555,8 @@ Ltac2 Notation "firstorder" tac(opt(thunk(tactic))) refs(opt(seq("using", list1(reference, ",")))) ids(opt(seq("with", list1(ident)))) := firstorder0 tac refs ids. + +(** now *) + +Ltac2 now0 t := t (); ltac1:(easy). +Ltac2 Notation "now" t(thunk(self)) := now0 t. -- cgit v1.2.3 From 62606e17ff4afe6a897607d45471b7f4d3ef54b8 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 2 Nov 2017 11:41:49 +0100 Subject: Binding the specialize tactic. --- src/tac2stdlib.ml | 6 ++++++ src/tac2tactics.ml | 5 +++++ src/tac2tactics.mli | 2 ++ tests/example2.v | 13 +++++++++++++ theories/Notations.v | 6 ++++++ theories/Std.v | 2 ++ 6 files changed, 34 insertions(+) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 60b6e70d58..6026b5b319 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -569,6 +569,12 @@ let () = define_prim3 "tac_constructorn" begin fun ev n bnd -> Tac2tactics.constructor_tac ev None n bnd end +let () = define_prim2 "tac_specialize" begin fun c ipat -> + let c = to_constr_with_bindings c in + let ipat = Value.to_option to_intro_pattern ipat in + Tac2tactics.specialize c ipat +end + let () = define_prim1 "tac_symmetry" begin fun cl -> let cl = to_clause cl in Tac2tactics.symmetry cl diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index f7ad057e86..eec0d2ab45 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -150,6 +150,11 @@ let split_with_bindings ev bnd = let bnd = mk_bindings bnd in Tactics.split_with_bindings ev [bnd] +let specialize c pat = + let c = mk_with_bindings c in + let pat = Option.map mk_intro_pattern pat in + Tactics.specialize c pat + let change pat c cl = let open Tac2ffi in Proofview.Goal.enter begin fun gl -> diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli index 7a4624ba2c..96c7b9214c 100644 --- a/src/tac2tactics.mli +++ b/src/tac2tactics.mli @@ -38,6 +38,8 @@ val left_with_bindings : evars_flag -> bindings -> unit tactic val right_with_bindings : evars_flag -> bindings -> unit tactic val split_with_bindings : evars_flag -> bindings -> unit tactic +val specialize : constr_with_bindings -> intro_pattern option -> unit tactic + val change : Pattern.constr_pattern option -> (constr array, constr) Tac2ffi.fun1 -> clause -> unit tactic val rewrite : diff --git a/tests/example2.v b/tests/example2.v index 46e4e43ed0..c953d25061 100644 --- a/tests/example2.v +++ b/tests/example2.v @@ -266,3 +266,16 @@ Proof. change (?a + 1 = 2) with (2 = $a + 1). reflexivity. Qed. + +Goal (forall n, n = 0 -> False) -> False. +Proof. +intros H. +specialize (H 0 eq_refl). +destruct H. +Qed. + +Goal (forall n, n = 0 -> False) -> False. +Proof. +intros H. +specialize (H 0 eq_refl) as []. +Qed. diff --git a/theories/Notations.v b/theories/Notations.v index dc09812254..f16a3a9161 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -204,6 +204,12 @@ Ltac2 Notation "econstructor" := Control.enter (fun () => Std.constructor true). Ltac2 Notation econstructor := econstructor. Ltac2 Notation "econstructor" n(tactic) bnd(thunk(with_bindings)) := constructor0 true n bnd. +Ltac2 specialize0 c pat := + enter_h false (fun _ c => Std.specialize c pat) c. + +Ltac2 Notation "specialize" c(thunk(seq(constr, with_bindings))) ipat(opt(seq("as", intropattern))) := + specialize0 c ipat. + Ltac2 elim0 ev c bnd use := let f ev ((c, bnd, use)) := Std.elim ev (c, bnd) use in enter_h ev f (fun () => c (), bnd (), use ()). diff --git a/theories/Std.v b/theories/Std.v index 389299f266..73b2ba02c4 100644 --- a/theories/Std.v +++ b/theories/Std.v @@ -227,6 +227,8 @@ Ltac2 @ external move : ident -> move_location -> unit := "ltac2" "tac_move". Ltac2 @ external intro : ident option -> move_location option -> unit := "ltac2" "tac_intro". +Ltac2 @ external specialize : constr_with_bindings -> intro_pattern option -> unit := "ltac2" "tac_specialize". + (** extratactics *) Ltac2 @ external discriminate : evar_flag -> destruction_arg option -> unit := "ltac2" "tac_discriminate". -- cgit v1.2.3 From 7e7964ddcc41363151d95cddd1a68b3dc70bb070 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 2 Nov 2017 15:56:51 +0100 Subject: Moving pattern type constraints to pattern AST. --- src/g_ltac2.ml4 | 16 ++++++++++------ src/tac2entries.ml | 14 ++++++++------ src/tac2expr.mli | 5 +++-- src/tac2intern.ml | 53 +++++++++++++++++++++++++++++++++++++---------------- src/tac2quote.ml | 16 +++++++++------- 5 files changed, 67 insertions(+), 37 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 557e2bcb9a..fca1b3045c 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -220,15 +220,15 @@ GEXTEND Gram | None -> te | Some args -> Loc.tag ~loc:!@loc @@ CTacFun (args, te) in - (pat, None, te) + (pat, te) ] ] ; let_binder: [ [ pats = LIST1 input_fun -> match pats with - | [(_, CPatVar _) as pat, None] -> (pat, None) - | ((_, CPatVar (Name id)) as pat, None) :: args -> (pat, Some args) - | [pat, None] -> (pat, None) + | [(_, CPatVar _) as pat] -> (pat, None) + | ((_, CPatVar (Name id)) as pat) :: args -> (pat, Some args) + | [pat] -> (pat, None) | _ -> CErrors.user_err ~loc:!@loc (str "Invalid pattern") ] ] ; @@ -257,8 +257,12 @@ GEXTEND Gram | l = Prim.ident -> Loc.tag ~loc:!@loc (Name l) ] ] ; input_fun: - [ [ b = tac2pat LEVEL "0" -> (b, None) - | "("; b = tac2pat; t = OPT [ ":"; t = tac2type -> t ]; ")" -> (b, t) ] ] + [ [ b = tac2pat LEVEL "0" -> b + | "("; b = tac2pat; t = OPT [ ":"; t = tac2type -> t ]; ")" -> + match t with + | None -> b + | Some t -> Loc.tag ~loc:!@loc @@ CPatCnv (b, t) + ] ] ; tac2def_body: [ [ name = binder; it = LIST0 input_fun; ":="; e = tac2expr -> diff --git a/src/tac2entries.ml b/src/tac2entries.ml index cefb4b13b8..e48bf02321 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -277,11 +277,15 @@ let fresh_var avoid x = in Namegen.next_ident_away_from (Id.of_string x) bad +let extract_pattern_type (loc, p as pat) = match p with +| CPatCnv (pat, ty) -> pat, Some ty +| CPatVar _ | CPatRef _ -> pat, None + (** Mangle recursive tactics *) let inline_rec_tactic tactics = let avoid = List.fold_left (fun accu ((_, id), _) -> Id.Set.add id accu) Id.Set.empty tactics in let map (id, e) = match snd e with - | CTacFun (pat, _) -> (id, pat, e) + | CTacFun (pat, _) -> (id, List.map extract_pattern_type pat, e) | _ -> let loc, _ = id in user_err ?loc (str "Recursive tactic definitions must be functions") @@ -295,11 +299,9 @@ let inline_rec_tactic tactics = in (** Fresh variables to abstract over the function patterns *) let _, vars = List.fold_left fold_var (avoid, []) pat in - let map_body ((loc, id), _, e) = (Loc.tag ?loc @@ CPatVar (Name id)), None, e in + let map_body ((loc, id), _, e) = (Loc.tag ?loc @@ CPatVar (Name id)), e in let bnd = List.map map_body tactics in - let pat_of_id (loc, id) = - ((Loc.tag ?loc @@ CPatVar (Name id)), None) - in + let pat_of_id (loc, id) = (Loc.tag ?loc @@ CPatVar (Name id)) in let var_of_id (loc, id) = let qid = (loc, qualid_of_ident id) in Loc.tag ?loc @@ CTacRef (RelId qid) @@ -590,7 +592,7 @@ let perform_notation syn st = let mk loc args = let map (na, e) = let loc = loc_of_tacexpr e in - ((Loc.tag ?loc @@ CPatVar na), None, e) + ((Loc.tag ?loc @@ CPatVar na), e) in let bnd = List.map map args in Loc.tag ~loc @@ CTacLet (false, bnd, syn.synext_exp) diff --git a/src/tac2expr.mli b/src/tac2expr.mli index 89152dffe7..60f10d360f 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -91,6 +91,7 @@ type atom = type raw_patexpr_r = | CPatVar of Name.t | CPatRef of ltac_constructor or_tuple or_relid * raw_patexpr list +| CPatCnv of raw_patexpr * raw_typexpr and raw_patexpr = raw_patexpr_r located @@ -98,9 +99,9 @@ type raw_tacexpr_r = | CTacAtm of atom | CTacRef of tacref or_relid | CTacCst of ltac_constructor or_tuple or_relid -| CTacFun of (raw_patexpr * raw_typexpr option) list * raw_tacexpr +| CTacFun of raw_patexpr list * raw_tacexpr | CTacApp of raw_tacexpr * raw_tacexpr list -| CTacLet of rec_flag * (raw_patexpr * raw_typexpr option * raw_tacexpr) list * raw_tacexpr +| CTacLet of rec_flag * (raw_patexpr * raw_tacexpr) list * raw_tacexpr | CTacCnv of raw_tacexpr * raw_typexpr | CTacSeq of raw_tacexpr * raw_tacexpr | CTacCse of raw_tacexpr * raw_taccase list diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 7b35cd55aa..9afdb3aedc 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -556,11 +556,13 @@ type glb_patexpr = | GPatVar of Name.t | GPatRef of ltac_constructor or_tuple * glb_patexpr list -let rec intern_patexpr env (_, pat) = match pat with +let rec intern_patexpr env (loc, pat) = match pat with | CPatVar na -> GPatVar na | CPatRef (qid, pl) -> let kn = get_constructor env qid in GPatRef (kn, List.map (fun p -> intern_patexpr env p) pl) +| CPatCnv (pat, ty) -> + user_err ?loc (str "Pattern not handled yet") type pattern_kind = | PKind_empty @@ -604,11 +606,16 @@ let rec ids_of_pattern accu (_, pat) = match pat with | CPatVar (Name id) -> Id.Set.add id accu | CPatRef (_, pl) -> List.fold_left ids_of_pattern accu pl +| CPatCnv (pat, _) -> ids_of_pattern accu pat let loc_of_relid = function | RelId (loc, _) -> loc | AbsKn _ -> None +let extract_pattern_type (loc, p as pat) = match p with +| CPatCnv (pat, ty) -> pat, Some ty +| CPatVar _ | CPatRef _ -> pat, None + (** Expand pattern: [p => t] becomes [x => match x with p => t end] *) let expand_pattern avoid bnd = let fold (avoid, bnd) (pat, t) = @@ -667,6 +674,7 @@ let rec intern_rec env (loc, e) = match e with let kn = get_constructor env qid in intern_constructor env loc kn [] | CTacFun (bnd, e) -> + let bnd = List.map extract_pattern_type bnd in let map (_, t) = match t with | None -> GTypVar (fresh_id env) | Some t -> intern_type env t @@ -689,8 +697,9 @@ let rec intern_rec env (loc, e) = match e with let map arg = (** Thunk alias arguments *) let loc = loc_of_tacexpr arg in - let var = [Loc.tag ?loc @@ CPatVar Anonymous, Some (Loc.tag ?loc @@ CTypRef (AbsKn (Tuple 0), []))] in - Loc.tag ?loc @@ CTacFun (var, arg) + let t_unit = Loc.tag ?loc @@ CTypRef (AbsKn (Tuple 0), []) in + let var = Loc.tag ?loc @@ CPatCnv (Loc.tag ?loc @@ CPatVar Anonymous, t_unit) in + Loc.tag ?loc @@ CTacFun ([var], arg) in let args = List.map map args in intern_rec env (Loc.tag ?loc @@ CTacApp (e, args)) @@ -706,6 +715,11 @@ let rec intern_rec env (loc, e) = match e with let ret = unify_arrow ?loc env ft t in (GTacApp (f, args), ret) | CTacLet (is_rec, el, e) -> + let map (pat, e) = + let (pat, ty) = extract_pattern_type pat in + (pat, ty, e) + in + let el = List.map map el in let fold accu (pat, _, e) = let ids = ids_of_pattern Id.Set.empty pat in let common = Id.Set.inter ids accu in @@ -826,7 +840,7 @@ and intern_let_rec env loc ids el e = let (loc, pat) = pat in let na = match pat with | CPatVar na -> na - | CPatRef _ -> + | CPatRef _ | CPatCnv _ -> user_err ?loc (str "This kind of pattern is forbidden in let-rec bindings") in let id = fresh_id env in @@ -961,6 +975,8 @@ and intern_case env loc e pl = else warn_redundant_clause ?loc () in brT + | CPatCnv _ -> + user_err ?loc (str "Pattern not handled yet") in let () = unify ?loc:(loc_of_tacexpr br) env tbr ret in intern_branch rem @@ -1216,10 +1232,10 @@ let rec globalize ids (loc, er as e) = match er with let knc = get_constructor () qid in Loc.tag ?loc @@ CTacCst (AbsKn knc) | CTacFun (bnd, e) -> - let fold (pats, accu) (pat, t) = + let fold (pats, accu) pat = let accu = ids_of_pattern accu pat in let pat = globalize_pattern ids pat in - ((pat, t) :: pats, accu) + (pat :: pats, accu) in let bnd, ids = List.fold_left fold ([], ids) bnd in let bnd = List.rev bnd in @@ -1230,13 +1246,14 @@ let rec globalize ids (loc, er as e) = match er with let el = List.map (fun e -> globalize ids e) el in Loc.tag ?loc @@ CTacApp (e, el) | CTacLet (isrec, bnd, e) -> - let fold accu (pat, _, _) = ids_of_pattern accu pat in + let fold accu (pat, _) = ids_of_pattern accu pat in let ext = List.fold_left fold Id.Set.empty bnd in let eids = Id.Set.union ext ids in let e = globalize eids e in - let map (qid, t, e) = + let map (qid, e) = let ids = if isrec then eids else ids in - (qid, t, globalize ids e) + let qid = globalize_pattern ids qid in + (qid, globalize ids e) in let bnd = List.map map bnd in Loc.tag ?loc @@ CTacLet (isrec, bnd, e) @@ -1281,6 +1298,9 @@ and globalize_pattern ids (loc, pr as p) = match pr with let cst = AbsKn knc in let pl = List.map (fun p -> globalize_pattern ids p) pl in Loc.tag ?loc @@ CPatRef (cst, pl) +| CPatCnv (pat, ty) -> + let pat = globalize_pattern ids pat in + Loc.tag ?loc @@ CPatCnv (pat, ty) (** Kernel substitution *) @@ -1419,6 +1439,10 @@ let rec subst_rawpattern subst (loc, pr as p) = match pr with let pl' = List.smartmap (fun p -> subst_rawpattern subst p) pl in let c' = subst_or_relid subst c in if pl' == pl && c' == c then p else Loc.tag ?loc @@ CPatRef (c', pl') +| CPatCnv (pat, ty) -> + let pat' = subst_rawpattern subst pat in + let ty' = subst_rawtype subst ty in + if pat' == pat && ty' == ty then p else Loc.tag ?loc @@ CPatCnv (pat', ty') (** Used for notations *) let rec subst_rawexpr subst (loc, tr as t) = match tr with @@ -1430,10 +1454,7 @@ let rec subst_rawexpr subst (loc, tr as t) = match tr with let ref' = subst_or_relid subst ref in if ref' == ref then t else Loc.tag ?loc @@ CTacCst ref' | CTacFun (bnd, e) -> - let map (na, t as p) = - let t' = Option.smartmap (fun t -> subst_rawtype subst t) t in - if t' == t then p else (na, t') - in + let map pat = subst_rawpattern subst pat in let bnd' = List.smartmap map bnd in let e' = subst_rawexpr subst e in if bnd' == bnd && e' == e then t else Loc.tag ?loc @@ CTacFun (bnd', e') @@ -1442,10 +1463,10 @@ let rec subst_rawexpr subst (loc, tr as t) = match tr with let el' = List.smartmap (fun e -> subst_rawexpr subst e) el in if e' == e && el' == el then t else Loc.tag ?loc @@ CTacApp (e', el') | CTacLet (isrec, bnd, e) -> - let map (na, t, e as p) = - let t' = Option.smartmap (fun t -> subst_rawtype subst t) t in + let map (na, e as p) = + let na' = subst_rawpattern subst na in let e' = subst_rawexpr subst e in - if t' == t && e' == e then p else (na, t', e') + if na' == na && e' == e then p else (na', e') in let bnd' = List.smartmap map bnd in let e' = subst_rawexpr subst e in diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 399c1199bd..33c4a97de1 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -54,8 +54,10 @@ let std_proj ?loc name = let thunk e = let t_unit = coq_core "unit" in let loc = Tac2intern.loc_of_tacexpr e in - let var = [Loc.tag ?loc @@ CPatVar (Anonymous), Some (Loc.tag ?loc @@ CTypRef (AbsKn (Other t_unit), []))] in - Loc.tag ?loc @@ CTacFun (var, e) + let ty = Loc.tag ?loc @@ CTypRef (AbsKn (Other t_unit), []) in + let pat = Loc.tag ?loc @@ CPatVar (Anonymous) in + let pat = Loc.tag ?loc @@ CPatCnv (pat, ty) in + Loc.tag ?loc @@ CTacFun ([pat], e) let of_pair f g (loc, (e1, e2)) = Loc.tag ?loc @@ CTacApp (Loc.tag ?loc @@ CTacCst (AbsKn (Tuple 2)), [f e1; g e2]) @@ -238,14 +240,14 @@ let abstract_vars loc vars tac = let get = global_ref ?loc (kername array_prefix "get") in let args = [of_variable (loc, id0); of_int (loc, n)] in let e = Loc.tag ?loc @@ CTacApp (get, args) in - let accu = (Loc.tag ?loc @@ CPatVar na, None, e) :: accu in + let accu = (Loc.tag ?loc @@ CPatVar na, e) :: accu in (n + 1, accu) in let (_, bnd) = List.fold_left build_bindings (0, []) vars in let tac = Loc.tag ?loc @@ CTacLet (false, bnd, tac) in (Name id0, tac) in - Loc.tag ?loc @@ CTacFun ([Loc.tag ?loc @@ CPatVar na, None], tac) + Loc.tag ?loc @@ CTacFun ([Loc.tag ?loc @@ CPatVar na], tac) let of_pattern p = inj_wit ?loc:p.CAst.loc wit_pattern p @@ -253,7 +255,7 @@ let of_pattern p = let of_conversion (loc, c) = match c with | QConvert c -> let pat = of_option ?loc of_pattern None in - let c = Loc.tag ?loc @@ CTacFun ([Loc.tag ?loc @@ CPatVar Anonymous, None], of_constr c) in + let c = Loc.tag ?loc @@ CTacFun ([Loc.tag ?loc @@ CPatVar Anonymous], of_constr c) in of_tuple ?loc [pat; c] | QConvertWith (pat, c) -> let vars = pattern_vars pat in @@ -389,7 +391,7 @@ let of_constr_matching (loc, m) = let vars = Id.Set.elements vars in let vars = List.map (fun id -> Name id) vars in let e = abstract_vars loc vars tac in - let e = Loc.tag ?loc @@ CTacFun ([Loc.tag ?loc @@ CPatVar na, None], e) in + let e = Loc.tag ?loc @@ CTacFun ([Loc.tag ?loc @@ CPatVar na], e) in let pat = inj_wit ?loc:ploc wit_pattern pat in of_tuple [knd; pat; e] in @@ -433,7 +435,7 @@ let of_goal_matching (loc, gm) = in let map (loc, (pat, tac)) = let (pat, hyps, hctx, subst, cctx) = mk_gpat pat in - let tac = Loc.tag ?loc @@ CTacFun ([Loc.tag ?loc @@ CPatVar cctx, None], tac) in + let tac = Loc.tag ?loc @@ CTacFun ([Loc.tag ?loc @@ CPatVar cctx], tac) in let tac = abstract_vars loc subst tac in let tac = abstract_vars loc hctx tac in let tac = abstract_vars loc hyps tac in -- cgit v1.2.3 From 2d0336671971489f217d666afde6537295b8c44a Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 2 Nov 2017 16:33:13 +0100 Subject: Factorizing entries for patterns with type constraints. --- src/g_ltac2.ml4 | 23 ++++++++++++++--------- 1 file changed, 14 insertions(+), 9 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index fca1b3045c..080fba7103 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -127,9 +127,19 @@ GEXTEND Gram [ "_" -> Loc.tag ~loc:!@loc @@ CPatVar Anonymous | "()" -> Loc.tag ~loc:!@loc @@ CPatRef (AbsKn (Tuple 0), []) | id = Prim.qualid -> pattern_of_qualid ~loc:!@loc id - | "("; pl = LIST0 tac2pat LEVEL "1" SEP ","; ")" -> - Loc.tag ~loc:!@loc @@ CPatRef (AbsKn (Tuple (List.length pl)), pl) ] - ] + | "("; p = atomic_tac2pat; ")" -> p + ] ] + ; + atomic_tac2pat: + [ [ -> + Loc.tag ~loc:!@loc @@ CPatRef (AbsKn (Tuple 0), []) + | p = tac2pat; ":"; t = tac2type -> + Loc.tag ~loc:!@loc @@ CPatCnv (p, t) + | p = tac2pat; ","; pl = LIST0 tac2pat SEP "," -> + let pl = p :: pl in + Loc.tag ~loc:!@loc @@ CPatRef (AbsKn (Tuple (List.length pl)), pl) + | p = tac2pat -> p + ] ] ; tac2expr: [ "6" RIGHTA @@ -257,12 +267,7 @@ GEXTEND Gram | l = Prim.ident -> Loc.tag ~loc:!@loc (Name l) ] ] ; input_fun: - [ [ b = tac2pat LEVEL "0" -> b - | "("; b = tac2pat; t = OPT [ ":"; t = tac2type -> t ]; ")" -> - match t with - | None -> b - | Some t -> Loc.tag ~loc:!@loc @@ CPatCnv (b, t) - ] ] + [ [ b = tac2pat LEVEL "0" -> b ] ] ; tac2def_body: [ [ name = binder; it = LIST0 input_fun; ":="; e = tac2expr -> -- cgit v1.2.3 From 290e9585ac3b0f6ece3f1966457fef3811f88d10 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 2 Nov 2017 16:44:57 +0100 Subject: Fix the horrible syntax that used to be valid for tuple matching. --- theories/Notations.v | 14 +++++++------- theories/Pattern.v | 20 ++++++++++---------- 2 files changed, 17 insertions(+), 17 deletions(-) diff --git a/theories/Notations.v b/theories/Notations.v index f16a3a9161..f4621656d6 100644 --- a/theories/Notations.v +++ b/theories/Notations.v @@ -46,7 +46,7 @@ Ltac2 orelse t f := match Control.case t with | Err e => f e | Val ans => - let ((x, k)) := ans in + let (x, k) := ans in Control.plus (fun _ => x) k end. @@ -54,7 +54,7 @@ Ltac2 ifcatch t s f := match Control.case t with | Err e => f e | Val ans => - let ((x, k)) := ans in + let (x, k) := ans in Control.plus (fun _ => s x) (fun e => s (k e)) end. @@ -73,11 +73,11 @@ Ltac2 rec repeat0 (t : unit -> unit) := Ltac2 Notation repeat := repeat0. -Ltac2 dispatch0 t ((head, tail)) := +Ltac2 dispatch0 t (head, tail) := match tail with | None => Control.enter (fun _ => t (); Control.dispatch head) | Some tacs => - let ((def, rem)) := tacs in + let (def, rem) := tacs in Control.enter (fun _ => t (); Control.extend head def rem) end. @@ -211,7 +211,7 @@ Ltac2 Notation "specialize" c(thunk(seq(constr, with_bindings))) ipat(opt(seq("a specialize0 c ipat. Ltac2 elim0 ev c bnd use := - let f ev ((c, bnd, use)) := Std.elim ev (c, bnd) use in + let f ev (c, bnd, use) := Std.elim ev (c, bnd) use in enter_h ev f (fun () => c (), bnd (), use ()). Ltac2 Notation "elim" c(thunk(constr)) bnd(thunk(with_bindings)) @@ -242,7 +242,7 @@ match cl with end. Ltac2 pose0 ev p := - enter_h ev (fun ev ((na, p)) => Std.pose na p) p. + enter_h ev (fun ev (na, p) => Std.pose na p) p. Ltac2 Notation "pose" p(thunk(pose)) := pose0 false p. @@ -382,7 +382,7 @@ Ltac2 Notation "native_compute" pl(opt(seq(pattern, occurrences))) cl(opt(clause Ltac2 Notation native_compute := native_compute. Ltac2 change0 p cl := - let ((pat, c)) := p in + let (pat, c) := p in Std.change pat c (default_on_concl cl). Ltac2 Notation "change" c(conversion) cl(opt(clause)) := change0 c cl. diff --git a/theories/Pattern.v b/theories/Pattern.v index 2c93918aee..ff7776b682 100644 --- a/theories/Pattern.v +++ b/theories/Pattern.v @@ -73,7 +73,7 @@ Ltac2 lazy_match0 t pats := | [] => Control.zero Match_failure | p :: m => let next _ := interp m in - let ((knd, pat, f)) := p in + let (knd, pat, f) := p in let p := match knd with | MatchPattern => (fun _ => @@ -82,7 +82,7 @@ Ltac2 lazy_match0 t pats := fun _ => f context bind) | MatchContext => (fun _ => - let ((context, bind)) := matches_subterm_vect pat t in + let (context, bind) := matches_subterm_vect pat t in fun _ => f context bind) end in Control.plus p next @@ -94,7 +94,7 @@ Ltac2 multi_match0 t pats := | [] => Control.zero Match_failure | p :: m => let next _ := interp m in - let ((knd, pat, f)) := p in + let (knd, pat, f) := p in let p := match knd with | MatchPattern => (fun _ => @@ -103,7 +103,7 @@ Ltac2 multi_match0 t pats := f context bind) | MatchContext => (fun _ => - let ((context, bind)) := matches_subterm_vect pat t in + let (context, bind) := matches_subterm_vect pat t in f context bind) end in Control.plus p next @@ -117,10 +117,10 @@ Ltac2 lazy_goal_match0 rev pats := | [] => Control.zero Match_failure | p :: m => let next _ := interp m in - let ((pat, f)) := p in - let ((phyps, pconcl)) := pat in + let (pat, f) := p in + let (phyps, pconcl) := pat in let cur _ := - let ((hids, hctx, subst, cctx)) := matches_goal rev phyps pconcl in + let (hids, hctx, subst, cctx) := matches_goal rev phyps pconcl in fun _ => f hids hctx subst cctx in Control.plus cur next @@ -132,10 +132,10 @@ Ltac2 multi_goal_match0 rev pats := | [] => Control.zero Match_failure | p :: m => let next _ := interp m in - let ((pat, f)) := p in - let ((phyps, pconcl)) := pat in + let (pat, f) := p in + let (phyps, pconcl) := pat in let cur _ := - let ((hids, hctx, subst, cctx)) := matches_goal rev phyps pconcl in + let (hids, hctx, subst, cctx) := matches_goal rev phyps pconcl in f hids hctx subst cctx in Control.plus cur next -- cgit v1.2.3 From 57f479f13c869408f51fbebc744c3b67a07d7f7c Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Sat, 4 Nov 2017 23:16:26 +0100 Subject: A possible fix after PR#6047 (a generic printer for ltac values). --- src/tac2core.ml | 4 ++-- src/tac2env.ml | 2 ++ 2 files changed, 4 insertions(+), 2 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index ce53b781f5..e476da7259 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -989,7 +989,7 @@ let () = let () = let pr_raw id = mt () in let pr_glb id = str "$" ++ Id.print id in - let pr_top _ = mt () in + let pr_top _ = Genprint.PrinterBasic mt in Genprint.register_print0 wit_ltac2_quotation pr_raw pr_glb pr_top (** Ltac2 in Ltac1 *) @@ -1017,7 +1017,7 @@ let () = let () = let pr_raw _ = mt () in let pr_glb e = Tac2print.pr_glbexpr e in - let pr_top _ = mt () in + let pr_top _ = Genprint.PrinterBasic mt in Genprint.register_print0 wit_ltac2 pr_raw pr_glb pr_top (** Built-in notation scopes *) diff --git a/src/tac2env.ml b/src/tac2env.ml index 2f1124c156..d0f286b396 100644 --- a/src/tac2env.ml +++ b/src/tac2env.ml @@ -280,6 +280,8 @@ let std_prefix = let wit_ltac2 = Genarg.make0 "ltac2:value" let wit_ltac2_quotation = Genarg.make0 "ltac2:quotation" +let () = Geninterp.register_val0 wit_ltac2 None +let () = Geninterp.register_val0 wit_ltac2_quotation None let is_constructor qid = let (_, id) = repr_qualid qid in -- cgit v1.2.3 From f773caf67f46bdaf80d9fd13f49b53c9a21cb091 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 6 Nov 2017 16:27:35 +0100 Subject: Generalize the use of repr in Tac2stdlib. --- _CoqProject | 2 + src/ltac2_plugin.mlpack | 1 + src/tac2extffi.ml | 40 +++++ src/tac2extffi.mli | 16 ++ src/tac2ffi.ml | 2 + src/tac2ffi.mli | 4 + src/tac2stdlib.ml | 428 ++++++++++++++++-------------------------------- src/tac2tactics.ml | 18 +- src/tac2tactics.mli | 16 +- 9 files changed, 226 insertions(+), 301 deletions(-) create mode 100644 src/tac2extffi.ml create mode 100644 src/tac2extffi.mli diff --git a/_CoqProject b/_CoqProject index df8cb00b76..eec66dc75e 100644 --- a/_CoqProject +++ b/_CoqProject @@ -25,6 +25,8 @@ src/tac2match.ml src/tac2match.mli src/tac2core.ml src/tac2core.mli +src/tac2extffi.ml +src/tac2extffi.mli src/tac2tactics.ml src/tac2tactics.mli src/tac2stdlib.ml diff --git a/src/ltac2_plugin.mlpack b/src/ltac2_plugin.mlpack index 40b91e4b53..2a25e825cb 100644 --- a/src/ltac2_plugin.mlpack +++ b/src/ltac2_plugin.mlpack @@ -8,6 +8,7 @@ Tac2entries Tac2quote Tac2match Tac2core +Tac2extffi Tac2tactics Tac2stdlib G_ltac2 diff --git a/src/tac2extffi.ml b/src/tac2extffi.ml new file mode 100644 index 0000000000..315c970f9e --- /dev/null +++ b/src/tac2extffi.ml @@ -0,0 +1,40 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* assert false) f + +(** More ML representations *) + +let to_qhyp v = match Value.to_block v with +| (0, [| i |]) -> AnonHyp (Value.to_int i) +| (1, [| id |]) -> NamedHyp (Value.to_ident id) +| _ -> assert false + +let qhyp = make_to_repr to_qhyp + +let to_bindings = function +| ValInt 0 -> NoBindings +| ValBlk (0, [| vl |]) -> + ImplicitBindings (Value.to_list Value.to_constr vl) +| ValBlk (1, [| vl |]) -> + ExplicitBindings ((Value.to_list (fun p -> to_pair to_qhyp Value.to_constr p) vl)) +| _ -> assert false + +let bindings = make_to_repr to_bindings + +let to_constr_with_bindings v = match Value.to_tuple v with +| [| c; bnd |] -> (Value.to_constr c, to_bindings bnd) +| _ -> assert false + +let constr_with_bindings = make_to_repr to_constr_with_bindings diff --git a/src/tac2extffi.mli b/src/tac2extffi.mli new file mode 100644 index 0000000000..f5251c3d0d --- /dev/null +++ b/src/tac2extffi.mli @@ -0,0 +1,16 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* 'a -> valexpr val repr_to : 'a repr -> valexpr -> 'a +val make_repr : ('a -> valexpr) -> (valexpr -> 'a) -> 'a repr + (** These functions allow to convert back and forth between OCaml and Ltac2 data representation. The [to_*] functions raise an anomaly whenever the data has not expected shape. *) @@ -109,6 +111,8 @@ val array : 'a repr -> 'a array repr val of_tuple : valexpr array -> valexpr val to_tuple : valexpr -> valexpr array +val of_pair : ('a -> valexpr) -> ('b -> valexpr) -> 'a * 'b -> valexpr +val to_pair : (valexpr -> 'a) -> (valexpr -> 'b) -> valexpr -> 'a * 'b val pair : 'a repr -> 'b repr -> ('a * 'b) repr val of_option : ('a -> valexpr) -> 'a option -> valexpr diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 6026b5b319..28d4967874 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -11,38 +11,25 @@ open Genredexpr open Tac2expr open Tac2ffi open Tac2types +open Tac2extffi open Proofview.Notations module Value = Tac2ffi +(** Make a representation with a dummy from function *) +let make_to_repr f = Tac2ffi.make_repr (fun _ -> assert false) f + let return x = Proofview.tclUNIT x let v_unit = Value.of_unit () -let thaw f = Tac2ffi.apply (Value.to_closure f) [v_unit] - -let to_pair f g v = match Value.to_tuple v with -| [| x; y |] -> (f x, g y) -| _ -> assert false +let thaw r f = Tac2ffi.app_fun1 f unit r () +let uthaw r f = Tac2ffi.app_fun1 (to_fun1 unit r f) unit r () +let thunk r = fun1 unit r let to_name c = match Value.to_option Value.to_ident c with | None -> Anonymous | Some id -> Name id -let to_qhyp v = match Value.to_block v with -| (0, [| i |]) -> AnonHyp (Value.to_int i) -| (1, [| id |]) -> NamedHyp (Value.to_ident id) -| _ -> assert false - -let to_bindings = function -| ValInt 0 -> NoBindings -| ValBlk (0, [| vl |]) -> - ImplicitBindings (Value.to_list Value.to_constr vl) -| ValBlk (1, [| vl |]) -> - ExplicitBindings ((Value.to_list (fun p -> to_pair to_qhyp Value.to_constr p) vl)) -| _ -> assert false - -let to_constr_with_bindings v = match Value.to_tuple v with -| [| c; bnd |] -> (Value.to_constr c, to_bindings bnd) -| _ -> assert false +let name = make_to_repr to_name let to_occurrences = function | ValInt 0 -> AllOccurrences @@ -51,6 +38,8 @@ let to_occurrences = function | ValBlk (1, [| vl |]) -> OnlyOccurrences (Value.to_list Value.to_int vl) | _ -> assert false +let occurrences = make_to_repr to_occurrences + let to_hyp_location_flag v = match Value.to_int v with | 0 -> InHyp | 1 -> InHypTypeOnly @@ -68,6 +57,8 @@ let to_clause v = match Value.to_tuple v with { onhyps = hyps; concl_occs = to_occurrences concl; } | _ -> assert false +let clause = make_to_repr to_clause + let to_red_flag v = match Value.to_tuple v with | [| beta; iota; fix; cofix; zeta; delta; const |] -> { @@ -81,11 +72,13 @@ let to_red_flag v = match Value.to_tuple v with } | _ -> assert false -let to_pattern_with_occs pat = - to_pair Value.to_pattern to_occurrences pat +let red_flags = make_to_repr to_red_flag + +let pattern_with_occs = pair pattern occurrences -let to_constr_with_occs c = - to_pair Value.to_constr to_occurrences c +let constr_with_occs = pair constr occurrences + +let reference_with_occs = pair reference occurrences let rec to_intro_pattern v = match Value.to_block v with | (0, [| b |]) -> IntroForthcoming (Value.to_bool b) @@ -121,14 +114,20 @@ and to_or_and_intro_pattern v = match Value.to_block v with and to_intro_patterns il = Value.to_list to_intro_pattern il +let intro_pattern = make_to_repr to_intro_pattern + +let intro_patterns = make_to_repr to_intro_patterns + let to_destruction_arg v = match Value.to_block v with | (0, [| c |]) -> - let c = thaw c >>= fun c -> return (to_constr_with_bindings c) in + let c = uthaw constr_with_bindings c in ElimOnConstr c | (1, [| id |]) -> ElimOnIdent (Value.to_ident id) | (2, [| n |]) -> ElimOnAnonHyp (Value.to_int n) | _ -> assert false +let destruction_arg = make_to_repr to_destruction_arg + let to_induction_clause v = match Value.to_tuple v with | [| arg; eqn; as_; in_ |] -> let arg = to_destruction_arg arg in @@ -139,6 +138,8 @@ let to_induction_clause v = match Value.to_tuple v with | _ -> assert false +let induction_clause = make_to_repr to_induction_clause + let to_assertion v = match Value.to_block v with | (0, [| ipat; t; tac |]) -> let to_tac t = Value.to_fun1 Value.unit Value.unit t in @@ -150,6 +151,8 @@ let to_assertion v = match Value.to_block v with AssertValue (Value.to_ident id, Value.to_constr c) | _ -> assert false +let assertion = make_to_repr to_assertion + let to_multi = function | ValBlk (0, [| n |]) -> Precisely (Value.to_int n) | ValBlk (1, [| n |]) -> UpTo (Value.to_int n) @@ -161,27 +164,35 @@ let to_rewriting v = match Value.to_tuple v with | [| orient; repeat; c |] -> let orient = Value.to_option Value.to_bool orient in let repeat = to_multi repeat in - let c = thaw c >>= fun c -> return (to_constr_with_bindings c) in + let c = uthaw constr_with_bindings c in (orient, repeat, c) | _ -> assert false +let rewriting = make_to_repr to_rewriting + let to_debug v = match Value.to_int v with | 0 -> Hints.Off | 1 -> Hints.Info | 2 -> Hints.Debug | _ -> assert false +let debug = make_to_repr to_debug + let to_strategy v = match Value.to_int v with | 0 -> Class_tactics.Bfs | 1 -> Class_tactics.Dfs | _ -> assert false +let strategy = make_to_repr to_strategy + let to_inversion_kind v = match Value.to_int v with | 0 -> Misctypes.SimpleInversion | 1 -> Misctypes.FullInversion | 2 -> Misctypes.FullInversionClear | _ -> assert false +let inversion_kind = make_to_repr to_inversion_kind + let to_move_location = function | ValInt 0 -> Misctypes.MoveFirst | ValInt 1 -> Misctypes.MoveLast @@ -189,6 +200,15 @@ let to_move_location = function | ValBlk (1, [|id|]) -> Misctypes.MoveBefore (Value.to_ident id) | _ -> assert false +let move_location = make_to_repr to_move_location + +let to_generalize_arg v = match Value.to_tuple v with +| [| c; occs; na |] -> + (Value.to_constr c, to_occurrences occs, to_name na) +| _ -> assert false + +let generalize_arg = make_to_repr to_generalize_arg + (** Standard tactics sharing their implementation with Ltac1 *) let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } @@ -199,189 +219,127 @@ let define_prim0 name tac = let tac _ = lift tac in Tac2env.define_primitive (pname name) (mk_closure arity_one tac) -let define_prim1 name tac = - let tac x = lift (tac x) in +let define_prim1 name r0 f = + let tac x = lift (f (Value.repr_to r0 x)) in Tac2env.define_primitive (pname name) (mk_closure arity_one tac) -let define_prim2 name tac = - let tac x y = lift (tac x y) in +let define_prim2 name r0 r1 f = + let tac x y = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y)) in Tac2env.define_primitive (pname name) (mk_closure (arity_suc arity_one) tac) -let define_prim3 name tac = - let tac x y z = lift (tac x y z) in +let define_prim3 name r0 r1 r2 f = + let tac x y z = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y) (Value.repr_to r2 z)) in Tac2env.define_primitive (pname name) (mk_closure (arity_suc (arity_suc arity_one)) tac) -let define_prim4 name tac = - let tac x y z u = lift (tac x y z u) in +let define_prim4 name r0 r1 r2 r3 f = + let tac x y z u = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y) (Value.repr_to r2 z) (Value.repr_to r3 u)) in Tac2env.define_primitive (pname name) (mk_closure (arity_suc (arity_suc (arity_suc arity_one))) tac) -let define_prim5 name tac = - let tac x y z u v = lift (tac x y z u v) in +let define_prim5 name r0 r1 r2 r3 r4 f = + let tac x y z u v = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y) (Value.repr_to r2 z) (Value.repr_to r3 u) (Value.repr_to r4 v)) in Tac2env.define_primitive (pname name) (mk_closure (arity_suc (arity_suc (arity_suc (arity_suc arity_one)))) tac) (** Tactics from Tacexpr *) -let () = define_prim2 "tac_intros" begin fun ev ipat -> - let ev = Value.to_bool ev in - let ipat = to_intro_patterns ipat in +let () = define_prim2 "tac_intros" bool intro_patterns begin fun ev ipat -> Tac2tactics.intros_patterns ev ipat end -let () = define_prim4 "tac_apply" begin fun adv ev cb ipat -> - let adv = Value.to_bool adv in - let ev = Value.to_bool ev in - let map_cb c = thaw c >>= fun c -> return (to_constr_with_bindings c) in - let cb = Value.to_list map_cb cb in - let map p = Value.to_option to_intro_pattern p in - let map_ipat p = to_pair Value.to_ident map p in - let ipat = Value.to_option map_ipat ipat in +let () = define_prim4 "tac_apply" bool bool (list (thunk constr_with_bindings)) (option (pair ident (option intro_pattern))) begin fun adv ev cb ipat -> Tac2tactics.apply adv ev cb ipat end -let () = define_prim3 "tac_elim" begin fun ev c copt -> - let ev = Value.to_bool ev in - let c = to_constr_with_bindings c in - let copt = Value.to_option to_constr_with_bindings copt in +let () = define_prim3 "tac_elim" bool constr_with_bindings (option constr_with_bindings) begin fun ev c copt -> Tac2tactics.elim ev c copt end -let () = define_prim2 "tac_case" begin fun ev c -> - let ev = Value.to_bool ev in - let c = to_constr_with_bindings c in +let () = define_prim2 "tac_case" bool constr_with_bindings begin fun ev c -> Tac2tactics.general_case_analysis ev c end -let () = define_prim1 "tac_generalize" begin fun cl -> - let cast v = match Value.to_tuple v with - | [| c; occs; na |] -> - (Value.to_constr c, to_occurrences occs, to_name na) - | _ -> assert false - in - let cl = Value.to_list cast cl in +let () = define_prim1 "tac_generalize" (list generalize_arg) begin fun cl -> Tac2tactics.generalize cl end -let () = define_prim1 "tac_assert" begin fun ast -> - let ast = to_assertion ast in +let () = define_prim1 "tac_assert" assertion begin fun ast -> Tac2tactics.assert_ ast end -let () = define_prim3 "tac_enough" begin fun c tac ipat -> - let c = Value.to_constr c in - let of_tac t = Proofview.tclIGNORE (thaw t) in - let tac = Value.to_option (fun t -> Value.to_option of_tac t) tac in - let ipat = Value.to_option to_intro_pattern ipat in +let () = define_prim3 "tac_enough" constr (option (option (thunk unit))) (option intro_pattern) begin fun c tac ipat -> + let tac = Option.map (fun o -> Option.map (fun f -> thaw unit f) o) tac in Tac2tactics.forward false tac ipat c end -let () = define_prim2 "tac_pose" begin fun idopt c -> - let na = to_name idopt in - let c = Value.to_constr c in +let () = define_prim2 "tac_pose" name constr begin fun na c -> Tactics.letin_tac None na c None Locusops.nowhere end -let () = define_prim3 "tac_set" begin fun ev p cl -> - let ev = Value.to_bool ev in - let cl = to_clause cl in +let () = define_prim3 "tac_set" bool (thunk (pair name constr)) clause begin fun ev p cl -> Proofview.tclEVARMAP >>= fun sigma -> - thaw p >>= fun p -> - let (na, c) = to_pair to_name Value.to_constr p in + thaw (pair name constr) p >>= fun (na, c) -> Tac2tactics.letin_pat_tac ev None na (sigma, c) cl end -let () = define_prim5 "tac_remember" begin fun ev na c eqpat cl -> - let ev = Value.to_bool ev in - let na = to_name na in - let cl = to_clause cl in - let eqpat = Value.to_option to_intro_pattern eqpat in +let () = define_prim5 "tac_remember" bool name (thunk constr) (option intro_pattern) clause begin fun ev na c eqpat cl -> let eqpat = Option.default (IntroNaming IntroAnonymous) eqpat in match eqpat with | IntroNaming eqpat -> Proofview.tclEVARMAP >>= fun sigma -> - thaw c >>= fun c -> - let c = Value.to_constr c in + thaw constr c >>= fun c -> Tac2tactics.letin_pat_tac ev (Some (true, eqpat)) na (sigma, c) cl | _ -> Tacticals.New.tclZEROMSG (Pp.str "Invalid pattern for remember") end -let () = define_prim3 "tac_destruct" begin fun ev ic using -> - let ev = Value.to_bool ev in - let ic = Value.to_list to_induction_clause ic in - let using = Value.to_option to_constr_with_bindings using in +let () = define_prim3 "tac_destruct" bool (list induction_clause) (option constr_with_bindings) begin fun ev ic using -> Tac2tactics.induction_destruct false ev ic using end -let () = define_prim3 "tac_induction" begin fun ev ic using -> - let ev = Value.to_bool ev in - let ic = Value.to_list to_induction_clause ic in - let using = Value.to_option to_constr_with_bindings using in +let () = define_prim3 "tac_induction" bool (list induction_clause) (option constr_with_bindings) begin fun ev ic using -> Tac2tactics.induction_destruct true ev ic using end -let () = define_prim1 "tac_red" begin fun cl -> - let cl = to_clause cl in +let () = define_prim1 "tac_red" clause begin fun cl -> Tac2tactics.reduce (Red false) cl end -let () = define_prim1 "tac_hnf" begin fun cl -> - let cl = to_clause cl in +let () = define_prim1 "tac_hnf" clause begin fun cl -> Tac2tactics.reduce Hnf cl end -let () = define_prim3 "tac_simpl" begin fun flags where cl -> - let flags = to_red_flag flags in - let where = Value.to_option to_pattern_with_occs where in - let cl = to_clause cl in +let () = define_prim3 "tac_simpl" red_flags (option pattern_with_occs) clause begin fun flags where cl -> Tac2tactics.simpl flags where cl end -let () = define_prim2 "tac_cbv" begin fun flags cl -> - let flags = to_red_flag flags in - let cl = to_clause cl in +let () = define_prim2 "tac_cbv" red_flags clause begin fun flags cl -> Tac2tactics.cbv flags cl end -let () = define_prim2 "tac_cbn" begin fun flags cl -> - let flags = to_red_flag flags in - let cl = to_clause cl in +let () = define_prim2 "tac_cbn" red_flags clause begin fun flags cl -> Tac2tactics.cbn flags cl end -let () = define_prim2 "tac_lazy" begin fun flags cl -> - let flags = to_red_flag flags in - let cl = to_clause cl in +let () = define_prim2 "tac_lazy" red_flags clause begin fun flags cl -> Tac2tactics.lazy_ flags cl end -let () = define_prim2 "tac_unfold" begin fun refs cl -> - let map v = to_pair Value.to_reference to_occurrences v in - let refs = Value.to_list map refs in - let cl = to_clause cl in +let () = define_prim2 "tac_unfold" (list reference_with_occs) clause begin fun refs cl -> Tac2tactics.unfold refs cl end -let () = define_prim2 "tac_fold" begin fun args cl -> - let args = Value.to_list Value.to_constr args in - let cl = to_clause cl in +let () = define_prim2 "tac_fold" (list constr) clause begin fun args cl -> Tac2tactics.reduce (Fold args) cl end -let () = define_prim2 "tac_pattern" begin fun where cl -> - let where = Value.to_list to_constr_with_occs where in - let cl = to_clause cl in +let () = define_prim2 "tac_pattern" (list constr_with_occs) clause begin fun where cl -> Tac2tactics.pattern where cl end -let () = define_prim2 "tac_vm" begin fun where cl -> - let where = Value.to_option to_pattern_with_occs where in - let cl = to_clause cl in +let () = define_prim2 "tac_vm" (option pattern_with_occs) clause begin fun where cl -> Tac2tactics.vm where cl end -let () = define_prim2 "tac_native" begin fun where cl -> - let where = Value.to_option to_pattern_with_occs where in - let cl = to_clause cl in +let () = define_prim2 "tac_native" (option pattern_with_occs) clause begin fun where cl -> Tac2tactics.native where cl end @@ -389,105 +347,71 @@ end let lift tac = tac >>= fun c -> Proofview.tclUNIT (Value.of_constr c) -let define_red1 name tac = - let tac x = lift (tac x) in +let define_red1 name r0 f = + let tac x = lift (f (Value.repr_to r0 x)) in Tac2env.define_primitive (pname name) (mk_closure arity_one tac) -let define_red2 name tac = - let tac x y = lift (tac x y) in +let define_red2 name r0 r1 f = + let tac x y = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y)) in Tac2env.define_primitive (pname name) (mk_closure (arity_suc arity_one) tac) -let define_red3 name tac = - let tac x y z = lift (tac x y z) in +let define_red3 name r0 r1 r2 f = + let tac x y z = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y) (Value.repr_to r2 z)) in Tac2env.define_primitive (pname name) (mk_closure (arity_suc (arity_suc arity_one)) tac) -let () = define_red1 "eval_red" begin fun c -> - let c = Value.to_constr c in +let () = define_red1 "eval_red" constr begin fun c -> Tac2tactics.eval_red c end -let () = define_red1 "eval_hnf" begin fun c -> - let c = Value.to_constr c in +let () = define_red1 "eval_hnf" constr begin fun c -> Tac2tactics.eval_hnf c end -let () = define_red3 "eval_simpl" begin fun flags where c -> - let flags = to_red_flag flags in - let where = Value.to_option to_pattern_with_occs where in - let c = Value.to_constr c in +let () = define_red3 "eval_simpl" red_flags (option pattern_with_occs) constr begin fun flags where c -> Tac2tactics.eval_simpl flags where c end -let () = define_red2 "eval_cbv" begin fun flags c -> - let flags = to_red_flag flags in - let c = Value.to_constr c in +let () = define_red2 "eval_cbv" red_flags constr begin fun flags c -> Tac2tactics.eval_cbv flags c end -let () = define_red2 "eval_cbn" begin fun flags c -> - let flags = to_red_flag flags in - let c = Value.to_constr c in +let () = define_red2 "eval_cbn" red_flags constr begin fun flags c -> Tac2tactics.eval_cbn flags c end -let () = define_red2 "eval_lazy" begin fun flags c -> - let flags = to_red_flag flags in - let c = Value.to_constr c in +let () = define_red2 "eval_lazy" red_flags constr begin fun flags c -> Tac2tactics.eval_lazy flags c end -let () = define_red2 "eval_unfold" begin fun refs c -> - let map v = to_pair Value.to_reference to_occurrences v in - let refs = Value.to_list map refs in - let c = Value.to_constr c in +let () = define_red2 "eval_unfold" (list reference_with_occs) constr begin fun refs c -> Tac2tactics.eval_unfold refs c end -let () = define_red2 "eval_fold" begin fun args c -> - let args = Value.to_list Value.to_constr args in - let c = Value.to_constr c in +let () = define_red2 "eval_fold" (list constr) constr begin fun args c -> Tac2tactics.eval_fold args c end -let () = define_red2 "eval_pattern" begin fun where c -> - let where = Value.to_list (fun p -> to_pair Value.to_constr to_occurrences p) where in - let c = Value.to_constr c in +let () = define_red2 "eval_pattern" (list constr_with_occs) constr begin fun where c -> Tac2tactics.eval_pattern where c end -let () = define_red2 "eval_vm" begin fun where c -> - let where = Value.to_option to_pattern_with_occs where in - let c = Value.to_constr c in +let () = define_red2 "eval_vm" (option pattern_with_occs) constr begin fun where c -> Tac2tactics.eval_vm where c end -let () = define_red2 "eval_native" begin fun where c -> - let where = Value.to_option to_pattern_with_occs where in - let c = Value.to_constr c in +let () = define_red2 "eval_native" (option pattern_with_occs) constr begin fun where c -> Tac2tactics.eval_native where c end -let () = define_prim3 "tac_change" begin fun pat c cl -> - let pat = Value.to_option (fun p -> Value.to_pattern p) pat in - let c = Value.to_fun1 (array constr) constr c in - let cl = to_clause cl in +let () = define_prim3 "tac_change" (option pattern) (fun1 (array constr) constr) clause begin fun pat c cl -> Tac2tactics.change pat c cl end -let () = define_prim4 "tac_rewrite" begin fun ev rw cl by -> - let ev = Value.to_bool ev in - let rw = Value.to_list to_rewriting rw in - let cl = to_clause cl in - let to_tac t = Proofview.tclIGNORE (thaw t) in - let by = Value.to_option to_tac by in +let () = define_prim4 "tac_rewrite" bool (list rewriting) clause (option (thunk unit)) begin fun ev rw cl by -> Tac2tactics.rewrite ev rw cl by end -let () = define_prim4 "tac_inversion" begin fun knd arg pat ids -> - let knd = to_inversion_kind knd in - let arg = to_destruction_arg arg in - let pat = Value.to_option to_intro_pattern pat in - let ids = Value.to_option (fun l -> Value.to_list Value.to_ident l) ids in +let () = define_prim4 "tac_inversion" inversion_kind destruction_arg (option intro_pattern) (option (list ident)) begin fun knd arg pat ids -> Tac2tactics.inversion knd arg pat ids end @@ -495,15 +419,11 @@ end let () = define_prim0 "tac_reflexivity" Tactics.intros_reflexivity -let () = define_prim2 "tac_move" begin fun id mv -> - let id = Value.to_ident id in - let mv = to_move_location mv in +let () = define_prim2 "tac_move" ident move_location begin fun id mv -> Tactics.move_hyp id mv end -let () = define_prim2 "tac_intro" begin fun id mv -> - let id = Value.to_option Value.to_ident id in - let mv = Value.to_option to_move_location mv in +let () = define_prim2 "tac_intro" (option ident) (option move_location) begin fun id mv -> let mv = Option.default Misctypes.MoveLast mv in Tactics.intro_move id mv end @@ -518,150 +438,112 @@ END let () = define_prim0 "tac_assumption" Tactics.assumption -let () = define_prim1 "tac_transitivity" begin fun c -> - let c = Value.to_constr c in +let () = define_prim1 "tac_transitivity" constr begin fun c -> Tactics.intros_transitivity (Some c) end let () = define_prim0 "tac_etransitivity" (Tactics.intros_transitivity None) -let () = define_prim1 "tac_cut" begin fun c -> - let c = Value.to_constr c in +let () = define_prim1 "tac_cut" constr begin fun c -> Tactics.cut c end -let () = define_prim2 "tac_left" begin fun ev bnd -> - let ev = Value.to_bool ev in - let bnd = to_bindings bnd in +let () = define_prim2 "tac_left" bool bindings begin fun ev bnd -> Tac2tactics.left_with_bindings ev bnd end -let () = define_prim2 "tac_right" begin fun ev bnd -> - let ev = Value.to_bool ev in - let bnd = to_bindings bnd in +let () = define_prim2 "tac_right" bool bindings begin fun ev bnd -> Tac2tactics.right_with_bindings ev bnd end -let () = define_prim1 "tac_introsuntil" begin fun h -> - Tactics.intros_until (to_qhyp h) +let () = define_prim1 "tac_introsuntil" qhyp begin fun h -> + Tactics.intros_until h end -let () = define_prim1 "tac_exactnocheck" begin fun c -> - Tactics.exact_no_check (Value.to_constr c) +let () = define_prim1 "tac_exactnocheck" constr begin fun c -> + Tactics.exact_no_check c end -let () = define_prim1 "tac_vmcastnocheck" begin fun c -> - Tactics.vm_cast_no_check (Value.to_constr c) +let () = define_prim1 "tac_vmcastnocheck" constr begin fun c -> + Tactics.vm_cast_no_check c end -let () = define_prim1 "tac_nativecastnocheck" begin fun c -> - Tactics.native_cast_no_check (Value.to_constr c) +let () = define_prim1 "tac_nativecastnocheck" constr begin fun c -> + Tactics.native_cast_no_check c end -let () = define_prim1 "tac_constructor" begin fun ev -> - let ev = Value.to_bool ev in +let () = define_prim1 "tac_constructor" bool begin fun ev -> Tactics.any_constructor ev None end -let () = define_prim3 "tac_constructorn" begin fun ev n bnd -> - let ev = Value.to_bool ev in - let n = Value.to_int n in - let bnd = to_bindings bnd in +let () = define_prim3 "tac_constructorn" bool int bindings begin fun ev n bnd -> Tac2tactics.constructor_tac ev None n bnd end -let () = define_prim2 "tac_specialize" begin fun c ipat -> - let c = to_constr_with_bindings c in - let ipat = Value.to_option to_intro_pattern ipat in +let () = define_prim2 "tac_specialize" constr_with_bindings (option intro_pattern) begin fun c ipat -> Tac2tactics.specialize c ipat end -let () = define_prim1 "tac_symmetry" begin fun cl -> - let cl = to_clause cl in +let () = define_prim1 "tac_symmetry" clause begin fun cl -> Tac2tactics.symmetry cl end -let () = define_prim2 "tac_split" begin fun ev bnd -> - let ev = Value.to_bool ev in - let bnd = to_bindings bnd in +let () = define_prim2 "tac_split" bool bindings begin fun ev bnd -> Tac2tactics.split_with_bindings ev bnd end -let () = define_prim1 "tac_rename" begin fun ids -> - let map c = match Value.to_tuple c with - | [|x; y|] -> (Value.to_ident x, Value.to_ident y) - | _ -> assert false - in - let ids = Value.to_list map ids in +let () = define_prim1 "tac_rename" (list (pair ident ident)) begin fun ids -> Tactics.rename_hyp ids end -let () = define_prim1 "tac_revert" begin fun ids -> - let ids = Value.to_list Value.to_ident ids in +let () = define_prim1 "tac_revert" (list ident) begin fun ids -> Tactics.revert ids end let () = define_prim0 "tac_admit" Proofview.give_up -let () = define_prim2 "tac_fix" begin fun idopt n -> - let idopt = Value.to_option Value.to_ident idopt in - let n = Value.to_int n in +let () = define_prim2 "tac_fix" (option ident) int begin fun idopt n -> Tactics.fix idopt n end -let () = define_prim1 "tac_cofix" begin fun idopt -> - let idopt = Value.to_option Value.to_ident idopt in +let () = define_prim1 "tac_cofix" (option ident) begin fun idopt -> Tactics.cofix idopt end -let () = define_prim1 "tac_clear" begin fun ids -> - let ids = Value.to_list Value.to_ident ids in +let () = define_prim1 "tac_clear" (list ident) begin fun ids -> Tactics.clear ids end -let () = define_prim1 "tac_keep" begin fun ids -> - let ids = Value.to_list Value.to_ident ids in +let () = define_prim1 "tac_keep" (list ident) begin fun ids -> Tactics.keep ids end -let () = define_prim1 "tac_clearbody" begin fun ids -> - let ids = Value.to_list Value.to_ident ids in +let () = define_prim1 "tac_clearbody" (list ident) begin fun ids -> Tactics.clear_body ids end (** Tactics from extratactics *) -let () = define_prim2 "tac_discriminate" begin fun ev arg -> - let ev = Value.to_bool ev in - let arg = Value.to_option to_destruction_arg arg in +let () = define_prim2 "tac_discriminate" bool (option destruction_arg) begin fun ev arg -> Tac2tactics.discriminate ev arg end -let () = define_prim3 "tac_injection" begin fun ev ipat arg -> - let ev = Value.to_bool ev in - let ipat = Value.to_option to_intro_patterns ipat in - let arg = Value.to_option to_destruction_arg arg in +let () = define_prim3 "tac_injection" bool (option intro_patterns) (option destruction_arg) begin fun ev ipat arg -> Tac2tactics.injection ev ipat arg end -let () = define_prim1 "tac_absurd" begin fun c -> - Contradiction.absurd (Value.to_constr c) +let () = define_prim1 "tac_absurd" constr begin fun c -> + Contradiction.absurd c end -let () = define_prim1 "tac_contradiction" begin fun c -> - let c = Value.to_option to_constr_with_bindings c in +let () = define_prim1 "tac_contradiction" (option constr_with_bindings) begin fun c -> Tac2tactics.contradiction c end -let () = define_prim4 "tac_autorewrite" begin fun all by ids cl -> - let all = Value.to_bool all in - let by = Value.to_option (fun tac -> Proofview.tclIGNORE (thaw tac)) by in - let ids = Value.to_list Value.to_ident ids in - let cl = to_clause cl in +let () = define_prim4 "tac_autorewrite" bool (option (thunk unit)) (list ident) clause begin fun all by ids cl -> Tac2tactics.autorewrite ~all by ids cl end -let () = define_prim1 "tac_subst" begin fun ids -> - let ids = Value.to_list Value.to_ident ids in +let () = define_prim1 "tac_subst" (list ident) begin fun ids -> Equality.subst ids end @@ -669,54 +551,28 @@ let () = define_prim0 "tac_substall" (return () >>= fun () -> Equality.subst_all (** Auto *) -let () = define_prim3 "tac_trivial" begin fun dbg lems dbs -> - let dbg = to_debug dbg in - let map c = thaw c >>= fun c -> return (Value.to_constr c) in - let lems = Value.to_list map lems in - let dbs = Value.to_option (fun l -> Value.to_list Value.to_ident l) dbs in +let () = define_prim3 "tac_trivial" debug (list (thunk constr)) (option (list ident)) begin fun dbg lems dbs -> Tac2tactics.trivial dbg lems dbs end -let () = define_prim5 "tac_eauto" begin fun dbg n p lems dbs -> - let dbg = to_debug dbg in - let n = Value.to_option Value.to_int n in - let p = Value.to_option Value.to_int p in - let map c = thaw c >>= fun c -> return (Value.to_constr c) in - let lems = Value.to_list map lems in - let dbs = Value.to_option (fun l -> Value.to_list Value.to_ident l) dbs in +let () = define_prim5 "tac_eauto" debug (option int) (option int) (list (thunk constr)) (option (list ident)) begin fun dbg n p lems dbs -> Tac2tactics.eauto dbg n p lems dbs end -let () = define_prim4 "tac_auto" begin fun dbg n lems dbs -> - let dbg = to_debug dbg in - let n = Value.to_option Value.to_int n in - let map c = thaw c >>= fun c -> return (Value.to_constr c) in - let lems = Value.to_list map lems in - let dbs = Value.to_option (fun l -> Value.to_list Value.to_ident l) dbs in +let () = define_prim4 "tac_auto" debug (option int) (list (thunk constr)) (option (list ident)) begin fun dbg n lems dbs -> Tac2tactics.auto dbg n lems dbs end -let () = define_prim4 "tac_newauto" begin fun dbg n lems dbs -> - let dbg = to_debug dbg in - let n = Value.to_option Value.to_int n in - let map c = thaw c >>= fun c -> return (Value.to_constr c) in - let lems = Value.to_list map lems in - let dbs = Value.to_option (fun l -> Value.to_list Value.to_ident l) dbs in +let () = define_prim4 "tac_newauto" debug (option int) (list (thunk constr)) (option (list ident)) begin fun dbg n lems dbs -> Tac2tactics.new_auto dbg n lems dbs end -let () = define_prim3 "tac_typeclasses_eauto" begin fun str n dbs -> - let str = Value.to_option to_strategy str in - let n = Value.to_option Value.to_int n in - let dbs = Value.to_option (fun l -> Value.to_list Value.to_ident l) dbs in +let () = define_prim3 "tac_typeclasses_eauto" (option strategy) (option int) (option (list ident)) begin fun str n dbs -> Tac2tactics.typeclasses_eauto str n dbs end (** Firstorder *) -let () = define_prim3 "tac_firstorder" begin fun tac refs ids -> - let tac = Value.to_option (fun t -> Proofview.tclIGNORE (thaw t)) tac in - let refs = Value.to_list Value.to_reference refs in - let ids = Value.to_list Value.to_ident ids in +let () = define_prim3 "tac_firstorder" (option (thunk unit)) (list reference) (list ident) begin fun tac refs ids -> Tac2tactics.firstorder tac refs ids end diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index eec0d2ab45..cd5709b130 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -11,6 +11,7 @@ open Util open Names open Globnames open Tac2types +open Tac2extffi open Genredexpr open Proofview.Notations @@ -92,7 +93,7 @@ let intros_patterns ev ipat = let apply adv ev cb cl = let map c = - let c = c >>= fun c -> return (mk_with_bindings c) in + let c = thaw constr_with_bindings c >>= fun p -> return (mk_with_bindings p) in None, Loc.tag (delayed_of_tactic c) in let cb = List.map map cb in @@ -174,7 +175,7 @@ let rewrite ev rw cl by = in let rw = List.map map_rw rw in let cl = mk_clause cl in - let by = Option.map (fun tac -> Tacticals.New.tclCOMPLETE tac, Equality.Naive) by in + let by = Option.map (fun tac -> Tacticals.New.tclCOMPLETE (thaw Tac2ffi.unit tac), Equality.Naive) by in Equality.general_multi_rewrite ev rw cl by let symmetry cl = @@ -369,23 +370,25 @@ let autorewrite ~all by ids cl = let cl = mk_clause cl in match by with | None -> Autorewrite.auto_multi_rewrite ?conds ids cl - | Some by -> Autorewrite.auto_multi_rewrite_with ?conds by ids cl + | Some by -> + let by = thaw Tac2ffi.unit by in + Autorewrite.auto_multi_rewrite_with ?conds by ids cl (** Auto *) let trivial debug lems dbs = - let lems = List.map delayed_of_tactic lems in + let lems = List.map (fun c -> delayed_of_thunk Tac2ffi.constr c) lems in let dbs = Option.map (fun l -> List.map Id.to_string l) dbs in Auto.h_trivial ~debug lems dbs let auto debug n lems dbs = - let lems = List.map delayed_of_tactic lems in + let lems = List.map (fun c -> delayed_of_thunk Tac2ffi.constr c) lems in let dbs = Option.map (fun l -> List.map Id.to_string l) dbs in Auto.h_auto ~debug n lems dbs let new_auto debug n lems dbs = let make_depth n = snd (Eauto.make_dimension n None) in - let lems = List.map delayed_of_tactic lems in + let lems = List.map (fun c -> delayed_of_thunk Tac2ffi.constr c) lems in match dbs with | None -> Auto.new_full_auto ~debug (make_depth n) lems | Some dbs -> @@ -393,7 +396,7 @@ let new_auto debug n lems dbs = Auto.new_auto ~debug (make_depth n) lems dbs let eauto debug n p lems dbs = - let lems = List.map delayed_of_tactic lems in + let lems = List.map (fun c -> delayed_of_thunk Tac2ffi.constr c) lems in let dbs = Option.map (fun l -> List.map Id.to_string l) dbs in Eauto.gen_eauto (Eauto.make_dimension n p) lems dbs @@ -447,6 +450,7 @@ let firstorder tac refs ids = let open Ground_plugin in (** FUCK YOU API *) let ids = List.map Id.to_string ids in + let tac = Option.map (fun tac -> thaw Tac2ffi.unit tac) tac in let tac : unit API.Proofview.tactic option = Obj.magic (tac : unit Proofview.tactic option) in let refs : API.Globnames.global_reference list = Obj.magic (refs : Globnames.global_reference list) in let ids : API.Hints.hint_db_name list = Obj.magic (ids : Hints.hint_db_name list) in diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli index 96c7b9214c..52e8a94c19 100644 --- a/src/tac2tactics.mli +++ b/src/tac2tactics.mli @@ -19,7 +19,7 @@ open Proofview val intros_patterns : evars_flag -> intro_pattern list -> unit tactic val apply : advanced_flag -> evars_flag -> - constr_with_bindings tactic list -> + constr_with_bindings thunk list -> (Id.t * intro_pattern option) option -> unit tactic val induction_destruct : rec_flag -> evars_flag -> @@ -43,7 +43,7 @@ val specialize : constr_with_bindings -> intro_pattern option -> unit tactic val change : Pattern.constr_pattern option -> (constr array, constr) Tac2ffi.fun1 -> clause -> unit tactic val rewrite : - evars_flag -> rewriting list -> clause -> unit tactic option -> unit tactic + evars_flag -> rewriting list -> clause -> unit thunk option -> unit tactic val symmetry : clause -> unit tactic @@ -101,18 +101,18 @@ val discriminate : evars_flag -> destruction_arg option -> unit tactic val injection : evars_flag -> intro_pattern list option -> destruction_arg option -> unit tactic -val autorewrite : all:bool -> unit tactic option -> Id.t list -> clause -> unit tactic +val autorewrite : all:bool -> unit thunk option -> Id.t list -> clause -> unit tactic -val trivial : Hints.debug -> constr tactic list -> Id.t list option -> +val trivial : Hints.debug -> constr thunk list -> Id.t list option -> unit Proofview.tactic -val auto : Hints.debug -> int option -> constr tactic list -> +val auto : Hints.debug -> int option -> constr thunk list -> Id.t list option -> unit Proofview.tactic -val new_auto : Hints.debug -> int option -> constr tactic list -> +val new_auto : Hints.debug -> int option -> constr thunk list -> Id.t list option -> unit Proofview.tactic -val eauto : Hints.debug -> int option -> int option -> constr tactic list -> +val eauto : Hints.debug -> int option -> int option -> constr thunk list -> Id.t list option -> unit Proofview.tactic val typeclasses_eauto : Class_tactics.search_strategy option -> int option -> @@ -122,4 +122,4 @@ val inversion : Misctypes.inversion_kind -> destruction_arg -> intro_pattern opt val contradiction : constr_with_bindings option -> unit tactic -val firstorder : unit Proofview.tactic option -> global_reference list -> Id.t list -> unit tactic +val firstorder : unit thunk option -> global_reference list -> Id.t list -> unit tactic -- cgit v1.2.3 From d0305718b4141aa08675743d4f85238301f37ad7 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Tue, 21 Nov 2017 14:15:37 +0100 Subject: [coq] Adapt to Coq's new functional EXTEND API. See https://github.com/coq/coq/pull/6197 --- src/g_ltac2.ml4 | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 080fba7103..c738cb65bd 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -831,10 +831,11 @@ let classify_ltac2 = function | StrSyn _ -> Vernacexpr.VtUnknown, Vernacexpr.VtNow | StrMut _ | StrVal _ | StrPrm _ | StrTyp _ | StrRun _ -> Vernac_classifier.classify_as_sideeff -VERNAC COMMAND EXTEND VernacDeclareTactic2Definition +VERNAC COMMAND FUNCTIONAL EXTEND VernacDeclareTactic2Definition | [ "Ltac2" ltac2_entry(e) ] => [ classify_ltac2 e ] -> [ - let local = Locality.LocalityFixme.consume () in - Tac2entries.register_struct ?local e + fun ~atts ~st -> let open Vernacinterp in + Tac2entries.register_struct ?local:atts.locality e; + st ] END -- cgit v1.2.3 From 2bd31e5fffbd6722f20016c3962088ab2008e2c0 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Fri, 24 Nov 2017 19:00:13 +0100 Subject: A possible fix after PR#6158 (raw/glob generic printers for ltac values). Can the printers exploit the ability to now take an environment? --- src/tac2core.ml | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index e476da7259..48cec86540 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -987,9 +987,9 @@ let () = Pretyping.register_constr_interp0 wit_ltac2_quotation interp let () = - let pr_raw id = mt () in - let pr_glb id = str "$" ++ Id.print id in - let pr_top _ = Genprint.PrinterBasic mt in + let pr_raw id = Genprint.PrinterBasic mt in + let pr_glb id = Genprint.PrinterBasic (fun () -> str "$" ++ Id.print id) in + let pr_top _ = Genprint.TopPrinterBasic mt in Genprint.register_print0 wit_ltac2_quotation pr_raw pr_glb pr_top (** Ltac2 in Ltac1 *) @@ -1015,9 +1015,9 @@ let () = Geninterp.register_interp0 wit_ltac2 interp let () = - let pr_raw _ = mt () in - let pr_glb e = Tac2print.pr_glbexpr e in - let pr_top _ = Genprint.PrinterBasic mt in + let pr_raw _ = Genprint.PrinterBasic mt in + let pr_glb e = Genprint.PrinterBasic (fun () -> Tac2print.pr_glbexpr e) in + let pr_top _ = Genprint.TopPrinterBasic mt in Genprint.register_print0 wit_ltac2 pr_raw pr_glb pr_top (** Built-in notation scopes *) -- cgit v1.2.3 From 3c2b1b7f99a1e06ad86a3c5dbf8369d773928e85 Mon Sep 17 00:00:00 2001 From: Théo Zimmermann Date: Wed, 15 Nov 2017 17:50:51 +0100 Subject: Adapt to removal of match_appsubterm. --- src/tac2core.ml | 4 ++-- src/tac2match.ml | 2 +- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 48cec86540..d21c1998da 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -561,7 +561,7 @@ let () = define2 "pattern_matches_subterm" pattern constr begin fun pat c -> Proofview.tclOR (return ans) (fun _ -> of_ans s) in pf_apply begin fun env sigma -> - let ans = Constr_matching.match_appsubterm env sigma pat c in + let ans = Constr_matching.match_subterm env sigma (Id.Set.empty,pat) c in of_ans ans end end @@ -593,7 +593,7 @@ let () = define2 "pattern_matches_subterm_vect" pattern constr begin fun pat c - Proofview.tclOR (return ans) (fun _ -> of_ans s) in pf_apply begin fun env sigma -> - let ans = Constr_matching.match_appsubterm env sigma pat c in + let ans = Constr_matching.match_subterm env sigma (Id.Set.empty,pat) c in of_ans ans end end diff --git a/src/tac2match.ml b/src/tac2match.ml index fef5647725..5035c9dba6 100644 --- a/src/tac2match.ml +++ b/src/tac2match.ml @@ -173,7 +173,7 @@ module PatternMatching (E:StaticEnvironment) = struct | Some nctx -> Proofview.tclOR (k (Some m_ctx) nctx) (fun e -> (map s e).stream k ctx) } in - map (Constr_matching.match_appsubterm E.env E.sigma p term) imatching_error + map (Constr_matching.match_subterm E.env E.sigma (Id.Set.empty,p) term) imatching_error let hyp_match_type pat hyps = pick hyps >>= fun decl -> -- cgit v1.2.3 From 46dfd18cd1744adbe9fe8463423c5a4484ebeb70 Mon Sep 17 00:00:00 2001 From: Enrico Tassi Date: Mon, 25 Dec 2017 21:17:48 +0100 Subject: adapt to API.mli removal --- _CoqProject | 1 - src/g_ltac2.ml4 | 3 +-- src/tac2core.ml | 8 ++------ src/tac2tactics.ml | 6 +----- 4 files changed, 4 insertions(+), 14 deletions(-) diff --git a/_CoqProject b/_CoqProject index eec66dc75e..5af42197ea 100644 --- a/_CoqProject +++ b/_CoqProject @@ -1,6 +1,5 @@ -R theories/ Ltac2 -I src/ --bypass-API src/tac2dyn.ml src/tac2dyn.mli diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index c738cb65bd..31eb6d9db5 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -91,8 +91,7 @@ let tac2def_mut = Gram.entry_create "tactic:tac2def_mut" let tac2def_run = Gram.entry_create "tactic:tac2def_run" let tac2mode = Gram.entry_create "vernac:ltac2_command" -(** FUCK YOU API *) -let ltac1_expr = (Obj.magic Pltac.tactic_expr : Tacexpr.raw_tactic_expr Gram.entry) +let ltac1_expr = Pltac.tactic_expr let inj_wit wit loc x = Loc.tag ~loc @@ CTacExt (wit, x) let inj_open_constr loc c = inj_wit Tac2quote.wit_open_constr loc c diff --git a/src/tac2core.ml b/src/tac2core.ml index d21c1998da..295b1b24ec 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -945,8 +945,7 @@ let () = let ist = { env_ist = Id.Map.empty } in let lfun = Tac2interp.set_env ist Id.Map.empty in let ist = Ltac_plugin.Tacinterp.default_ist () in - (** FUCK YOU API *) - let ist = { ist with API.Geninterp.lfun = (Obj.magic lfun) } in + let ist = { ist with Geninterp.lfun = lfun } in let tac = (Obj.magic Ltac_plugin.Tacinterp.eval_tactic_ist ist tac : unit Proofview.tactic) in let wrap (e, info) = set_bt info >>= fun info -> Proofview.tclZERO ~info e in Proofview.tclOR tac wrap >>= fun () -> @@ -995,8 +994,7 @@ let () = (** Ltac2 in Ltac1 *) let () = - (** FUCK YOU API *) - let e = (Obj.magic Tac2entries.Pltac.tac2expr : _ API.Pcoq.Gram.entry) in + let e = Tac2entries.Pltac.tac2expr in let inject (loc, v) = Tacexpr.TacGeneric (in_gen (rawwit wit_ltac2) v) in Ltac_plugin.Tacentries.create_ltac_quotation "ltac2" inject (e, None) @@ -1004,8 +1002,6 @@ let () = let open Ltac_plugin in let open Tacinterp in let idtac = Value.of_closure (default_ist ()) (Tacexpr.TacId []) in - (** FUCK YOU API *) - let idtac = (Obj.magic idtac : Geninterp.Val.t) in let interp ist tac = (* let ist = Tac2interp.get_env ist.Geninterp.lfun in *) let ist = { env_ist = Id.Map.empty } in diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index cd5709b130..e7d5578b6e 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -448,10 +448,6 @@ let contradiction c = let firstorder tac refs ids = let open Ground_plugin in - (** FUCK YOU API *) let ids = List.map Id.to_string ids in let tac = Option.map (fun tac -> thaw Tac2ffi.unit tac) tac in - let tac : unit API.Proofview.tactic option = Obj.magic (tac : unit Proofview.tactic option) in - let refs : API.Globnames.global_reference list = Obj.magic (refs : Globnames.global_reference list) in - let ids : API.Hints.hint_db_name list = Obj.magic (ids : Hints.hint_db_name list) in - (Obj.magic (G_ground.gen_ground_tac true tac refs ids : unit API.Proofview.tactic) : unit Proofview.tactic) + G_ground.gen_ground_tac true tac refs ids -- cgit v1.2.3 From 2e015ec8e2958c2848c33a152dd883e048069a7d Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Wed, 14 Feb 2018 06:24:02 +0100 Subject: [coq] Adapt to coq/coq#6745 Nothing remarkable. --- src/tac2intern.ml | 4 ++-- src/tac2qexpr.mli | 2 +- src/tac2quote.ml | 2 +- src/tac2tactics.ml | 6 +++--- 4 files changed, 7 insertions(+), 7 deletions(-) diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 9afdb3aedc..dc142043e8 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -505,7 +505,7 @@ let check_redundant_clause = function let get_variable0 mem var = match var with | RelId (loc, qid) -> let (dp, id) = repr_qualid qid in - if DirPath.is_empty dp && mem id then ArgVar (loc, id) + if DirPath.is_empty dp && mem id then ArgVar CAst.(make ?loc id) else let kn = try Tac2env.locate_ltac qid @@ -652,7 +652,7 @@ let rec intern_rec env (loc, e) = match e with | CTacAtm atm -> intern_atm env atm | CTacRef qid -> begin match get_variable env qid with - | ArgVar (_, id) -> + | ArgVar {CAst.v=id} -> let sch = Id.Map.find id env.env_var in (GTacVar id, fresh_mix_type_scheme env sch) | ArgArg (TacConstant kn) -> diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index ad52884ca6..2f6c97f08b 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -136,7 +136,7 @@ type constr_matching = constr_match_branch list located type goal_match_pattern_r = { q_goal_match_concl : constr_match_pattern; - q_goal_match_hyps : (Name.t located * constr_match_pattern) list; + q_goal_match_hyps : (Misctypes.lname * constr_match_pattern) list; } type goal_match_pattern = goal_match_pattern_r located diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 33c4a97de1..d0c1365eff 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -425,7 +425,7 @@ let of_goal_matching (loc, gm) = let map (_, _, pat, knd) = of_tuple [knd; of_pattern pat] in let concl = of_tuple [concl_knd; of_pattern concl_pat] in let r = of_tuple [of_list ?loc map hyps_pats; concl] in - let hyps = List.map (fun ((_, na), _, _, _) -> na) hyps_pats in + let hyps = List.map (fun ({CAst.v=na}, _, _, _) -> na) hyps_pats in let map (_, na, _, _) = na in let hctx = List.map map hyps_pats in (** Order of elements is crucial here! *) diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index e7d5578b6e..65cdef0f3f 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -107,7 +107,7 @@ let mk_destruction_arg = function | ElimOnConstr c -> let c = c >>= fun c -> return (mk_with_bindings c) in Misctypes.ElimOnConstr (delayed_of_tactic c) -| ElimOnIdent id -> Misctypes.ElimOnIdent (Loc.tag id) +| ElimOnIdent id -> Misctypes.ElimOnIdent CAst.(make id) | ElimOnAnonHyp n -> Misctypes.ElimOnAnonHyp n let mk_induction_clause (arg, eqn, as_, occ) = @@ -343,7 +343,7 @@ let on_destruction_arg tac ev arg = let flags = tactic_infer_flags ev in let (sigma', c) = Unification.finish_evar_resolution ~flags env sigma' (sigma, c) in Proofview.tclUNIT (Some sigma', Misctypes.ElimOnConstr (c, lbind)) - | ElimOnIdent id -> Proofview.tclUNIT (None, Misctypes.ElimOnIdent (Loc.tag id)) + | ElimOnIdent id -> Proofview.tclUNIT (None, Misctypes.ElimOnIdent CAst.(make id)) | ElimOnAnonHyp n -> Proofview.tclUNIT (None, Misctypes.ElimOnAnonHyp n) in arg >>= fun (sigma', arg) -> @@ -429,7 +429,7 @@ let inversion knd arg pat ids = | None -> assert false | Some (_, Misctypes.ElimOnAnonHyp n) -> Inv.inv_clause knd pat ids (AnonHyp n) - | Some (_, Misctypes.ElimOnIdent (_, id)) -> + | Some (_, Misctypes.ElimOnIdent {CAst.v=id}) -> Inv.inv_clause knd pat ids (NamedHyp id) | Some (_, Misctypes.ElimOnConstr c) -> let open Misctypes in -- cgit v1.2.3 From cf1d983cfd42ae4a7e1e01c6cab348fc51233c65 Mon Sep 17 00:00:00 2001 From: Enrico Tassi Date: Tue, 6 Feb 2018 18:19:15 +0100 Subject: adapt to Coq#6676 --- src/tac2core.ml | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 295b1b24ec..1afaea8bd9 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -514,9 +514,9 @@ let () = define3 "constr_in_context" ident constr closure begin fun id t c -> let (sigma, (evt, _)) = Evarutil.new_type_evar nenv sigma Evd.univ_flexible in let (sigma, evk) = Evarutil.new_pure_evar (Environ.named_context_val nenv) sigma evt in Proofview.Unsafe.tclEVARS sigma >>= fun () -> - Proofview.Unsafe.tclSETGOALS [evk] >>= fun () -> + Proofview.Unsafe.tclSETGOALS [Proofview.with_empty_state evk] >>= fun () -> thaw c >>= fun _ -> - Proofview.Unsafe.tclSETGOALS [Proofview.Goal.goal (Proofview.Goal.assume gl)] >>= fun () -> + Proofview.Unsafe.tclSETGOALS [Proofview.with_empty_state (Proofview.Goal.goal (Proofview.Goal.assume gl))] >>= fun () -> let args = List.map (fun d -> EConstr.mkVar (get_id d)) (EConstr.named_context env) in let args = Array.of_list (EConstr.mkRel 1 :: args) in let ans = EConstr.mkEvar (evk, args) in @@ -700,7 +700,7 @@ let () = define1 "new_goal" int begin fun ev -> let ev = Evar.unsafe_of_int ev in Proofview.tclEVARMAP >>= fun sigma -> if Evd.mem sigma ev then - Proofview.Unsafe.tclNEWGOALS [ev] <*> Proofview.tclUNIT v_unit + Proofview.Unsafe.tclNEWGOALS [Proofview.with_empty_state ev] <*> Proofview.tclUNIT v_unit else throw err_notfound end -- cgit v1.2.3 From 5cb74f82165e88c5e527ad757b007df1fbe5f1b3 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Wed, 14 Feb 2018 02:05:42 +0100 Subject: [api] Remove some deprecation warnings. Trivial commit. --- src/tac2core.ml | 4 +++- src/tac2ffi.mli | 2 +- src/tac2quote.ml | 2 +- 3 files changed, 5 insertions(+), 3 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 1afaea8bd9..c16e72b801 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -894,7 +894,9 @@ let () = let () = let intern self ist c = - let _, pat = Constrintern.intern_constr_pattern ist.Genintern.genv ~as_type:false c in + let env = ist.Genintern.genv in + let sigma = Evd.from_env env in + let _, pat = Constrintern.intern_constr_pattern env sigma ~as_type:false c in GlbVal pat, gtypref t_pattern in let print env pat = str "pattern:(" ++ Printer.pr_lconstr_pattern_env env Evd.empty pat ++ str ")" in diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index 4d77c989f3..1bf86b516a 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -165,7 +165,7 @@ val val_constant : Constant.t Val.tag val val_constructor : constructor Val.tag val val_projection : Projection.t Val.tag val val_case : Constr.case_info Val.tag -val val_univ : Univ.universe_level Val.tag +val val_univ : Univ.Level.t Val.tag val val_free : Id.Set.t Val.tag val val_exn : Exninfo.iexn Tac2dyn.Val.tag diff --git a/src/tac2quote.ml b/src/tac2quote.ml index d0c1365eff..829f13344c 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -221,7 +221,7 @@ let pattern_vars pat = let () = check_pattern_id ?loc:pat.CAst.loc id in Id.Set.add id accu | _ -> - Topconstr.fold_constr_expr_with_binders (fun _ () -> ()) aux () accu pat + Constrexpr_ops.fold_constr_expr_with_binders (fun _ () -> ()) aux () accu pat in aux () Id.Set.empty pat -- cgit v1.2.3 From ba27c2e21f80b41dfd837e6c6b15f82ca405cf04 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Wed, 14 Feb 2018 02:06:02 +0100 Subject: [coq] Adapt to coq/coq#6511 coq/coq#6511 contains EConstr-related changes. --- src/tac2entries.ml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/tac2entries.ml b/src/tac2entries.ml index e48bf02321..fa498ab44f 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -344,7 +344,7 @@ let register_ltac ?(local = false) ?(mut = false) isrec tactics = in let () = if exists then - user_err ?loc (str "Tactic " ++ Nameops.pr_id id ++ str " already exists") + user_err ?loc (str "Tactic " ++ Names.Id.print id ++ str " already exists") in (id, e, t) in @@ -539,7 +539,7 @@ let parse_scope = function if Id.Map.mem id !scope_table then Id.Map.find id !scope_table toks else - CErrors.user_err ?loc (str "Unknown scope" ++ spc () ++ Nameops.pr_id id) + CErrors.user_err ?loc (str "Unknown scope" ++ spc () ++ Names.Id.print id) | SexprStr (_, str) -> let v_unit = Loc.tag @@ CTacCst (AbsKn (Tuple 0)) in ScopeRule (Extend.Atoken (Tok.IDENT str), (fun _ -> v_unit)) -- cgit v1.2.3 From bfa24f9435d5ebeb9c5abb81b8ae0ef7e305f83f Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Mon, 5 Mar 2018 23:48:50 +0100 Subject: [coq] Adapt to correct LTAC module packing coq/coq#6869 --- src/tac2core.ml | 2 +- src/tac2quote.mli | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index c16e72b801..62fc9db292 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -997,7 +997,7 @@ let () = let () = let e = Tac2entries.Pltac.tac2expr in - let inject (loc, v) = Tacexpr.TacGeneric (in_gen (rawwit wit_ltac2) v) in + let inject (loc, v) = Ltac_plugin.Tacexpr.TacGeneric (in_gen (rawwit wit_ltac2) v) in Ltac_plugin.Tacentries.create_ltac_quotation "ltac2" inject (e, None) let () = diff --git a/src/tac2quote.mli b/src/tac2quote.mli index 3f6c9a55e5..1f22fe6c74 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -99,4 +99,4 @@ val wit_constr : (Constrexpr.constr_expr, Glob_term.glob_constr) Arg.tag val wit_open_constr : (Constrexpr.constr_expr, Glob_term.glob_constr) Arg.tag -val wit_ltac1 : (Tacexpr.raw_tactic_expr, Tacexpr.glob_tactic_expr) Arg.tag +val wit_ltac1 : (Ltac_plugin.Tacexpr.raw_tactic_expr, Ltac_plugin.Tacexpr.glob_tactic_expr) Arg.tag -- cgit v1.2.3 From d2d1fe30e3defa18dd966bf8160df81fc1e72e31 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 9 Mar 2018 10:27:39 -0300 Subject: Fix compilation after the change of API in options. --- src/tac2entries.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/tac2entries.ml b/src/tac2entries.ml index fa498ab44f..1631880c71 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -919,7 +919,7 @@ let load_ltac2_init _ (_, ()) = Hook.get f_register_constr_quotations () let open_ltac2_init _ (_, ()) = - Goptions.set_string_option_value_gen None ["Default"; "Proof"; "Mode"] "Ltac2" + Goptions.set_string_option_value_gen ["Default"; "Proof"; "Mode"] "Ltac2" (** Dummy object that register global rules when Require is called *) let inTac2Init : unit -> obj = -- cgit v1.2.3 From 123de18a0886233b047ef2bad4bd7b3694f2abcc Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Mon, 5 Mar 2018 22:10:45 +0100 Subject: [coq] Adapt to coq/coq#6831. This removes uses of `Loc.t` in favor of `CAst.t`. --- src/g_ltac2.ml4 | 302 ++++++++++++++++++++++++++-------------------------- src/tac2core.ml | 31 +++--- src/tac2entries.ml | 94 ++++++++-------- src/tac2entries.mli | 11 +- src/tac2expr.mli | 24 ++--- src/tac2intern.ml | 173 +++++++++++++++--------------- src/tac2intern.mli | 3 - src/tac2qexpr.mli | 99 +++++++++-------- src/tac2quote.ml | 129 +++++++++++----------- src/tac2quote.mli | 17 ++- src/tac2tactics.ml | 24 ++--- 11 files changed, 447 insertions(+), 460 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 31eb6d9db5..f4818f4ece 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -93,17 +93,17 @@ let tac2mode = Gram.entry_create "vernac:ltac2_command" let ltac1_expr = Pltac.tactic_expr -let inj_wit wit loc x = Loc.tag ~loc @@ CTacExt (wit, x) +let inj_wit wit loc x = CAst.make ~loc @@ CTacExt (wit, x) let inj_open_constr loc c = inj_wit Tac2quote.wit_open_constr loc c let inj_pattern loc c = inj_wit Tac2quote.wit_pattern loc c let inj_reference loc c = inj_wit Tac2quote.wit_reference loc c let inj_ltac1 loc e = inj_wit Tac2quote.wit_ltac1 loc e let pattern_of_qualid ?loc id = - if Tac2env.is_constructor (snd id) then Loc.tag ?loc @@ CPatRef (RelId id, []) + if Tac2env.is_constructor id.CAst.v then CAst.make ?loc @@ CPatRef (RelId id, []) else - let (dp, id) = Libnames.repr_qualid (snd id) in - if DirPath.is_empty dp then Loc.tag ?loc @@ CPatVar (Name id) + let (dp, id) = Libnames.repr_qualid id.CAst.v in + if DirPath.is_empty dp then CAst.make ?loc @@ CPatVar (Name id) else CErrors.user_err ?loc (Pp.str "Syntax error") @@ -113,73 +113,73 @@ GEXTEND Gram tac2pat: [ "1" LEFTA [ id = Prim.qualid; pl = LIST1 tac2pat LEVEL "0" -> - if Tac2env.is_constructor (snd id) then - Loc.tag ~loc:!@loc @@ CPatRef (RelId id, pl) + if Tac2env.is_constructor (id.CAst.v) then + CAst.make ~loc:!@loc @@ CPatRef (RelId id, pl) else CErrors.user_err ~loc:!@loc (Pp.str "Syntax error") | id = Prim.qualid -> pattern_of_qualid ~loc:!@loc id - | "["; "]" -> Loc.tag ~loc:!@loc @@ CPatRef (AbsKn (Other Tac2core.Core.c_nil), []) + | "["; "]" -> CAst.make ~loc:!@loc @@ CPatRef (AbsKn (Other Tac2core.Core.c_nil), []) | p1 = tac2pat; "::"; p2 = tac2pat -> - Loc.tag ~loc:!@loc @@ CPatRef (AbsKn (Other Tac2core.Core.c_cons), [p1; p2]) + CAst.make ~loc:!@loc @@ CPatRef (AbsKn (Other Tac2core.Core.c_cons), [p1; p2]) ] | "0" - [ "_" -> Loc.tag ~loc:!@loc @@ CPatVar Anonymous - | "()" -> Loc.tag ~loc:!@loc @@ CPatRef (AbsKn (Tuple 0), []) + [ "_" -> CAst.make ~loc:!@loc @@ CPatVar Anonymous + | "()" -> CAst.make ~loc:!@loc @@ CPatRef (AbsKn (Tuple 0), []) | id = Prim.qualid -> pattern_of_qualid ~loc:!@loc id | "("; p = atomic_tac2pat; ")" -> p ] ] ; atomic_tac2pat: [ [ -> - Loc.tag ~loc:!@loc @@ CPatRef (AbsKn (Tuple 0), []) + CAst.make ~loc:!@loc @@ CPatRef (AbsKn (Tuple 0), []) | p = tac2pat; ":"; t = tac2type -> - Loc.tag ~loc:!@loc @@ CPatCnv (p, t) + CAst.make ~loc:!@loc @@ CPatCnv (p, t) | p = tac2pat; ","; pl = LIST0 tac2pat SEP "," -> let pl = p :: pl in - Loc.tag ~loc:!@loc @@ CPatRef (AbsKn (Tuple (List.length pl)), pl) + CAst.make ~loc:!@loc @@ CPatRef (AbsKn (Tuple (List.length pl)), pl) | p = tac2pat -> p ] ] ; tac2expr: [ "6" RIGHTA - [ e1 = SELF; ";"; e2 = SELF -> Loc.tag ~loc:!@loc @@ CTacSeq (e1, e2) ] + [ e1 = SELF; ";"; e2 = SELF -> CAst.make ~loc:!@loc @@ CTacSeq (e1, e2) ] | "5" [ "fun"; it = LIST1 input_fun ; "=>"; body = tac2expr LEVEL "6" -> - Loc.tag ~loc:!@loc @@ CTacFun (it, body) + CAst.make ~loc:!@loc @@ CTacFun (it, body) | "let"; isrec = rec_flag; lc = LIST1 let_clause SEP "with"; "in"; e = tac2expr LEVEL "6" -> - Loc.tag ~loc:!@loc @@ CTacLet (isrec, lc, e) + CAst.make ~loc:!@loc @@ CTacLet (isrec, lc, e) | "match"; e = tac2expr LEVEL "5"; "with"; bl = branches; "end" -> - Loc.tag ~loc:!@loc @@ CTacCse (e, bl) + CAst.make ~loc:!@loc @@ CTacCse (e, bl) ] | "4" LEFTA [ ] | "::" RIGHTA [ e1 = tac2expr; "::"; e2 = tac2expr -> - Loc.tag ~loc:!@loc @@ CTacApp (Loc.tag ~loc:!@loc @@ CTacCst (AbsKn (Other Tac2core.Core.c_cons)), [e1; e2]) + CAst.make ~loc:!@loc @@ CTacApp (CAst.make ~loc:!@loc @@ CTacCst (AbsKn (Other Tac2core.Core.c_cons)), [e1; e2]) ] | [ e0 = SELF; ","; el = LIST1 NEXT SEP "," -> let el = e0 :: el in - Loc.tag ~loc:!@loc @@ CTacApp (Loc.tag ~loc:!@loc @@ CTacCst (AbsKn (Tuple (List.length el))), el) ] + CAst.make ~loc:!@loc @@ CTacApp (CAst.make ~loc:!@loc @@ CTacCst (AbsKn (Tuple (List.length el))), el) ] | "1" LEFTA [ e = tac2expr; el = LIST1 tac2expr LEVEL "0" -> - Loc.tag ~loc:!@loc @@ CTacApp (e, el) + CAst.make ~loc:!@loc @@ CTacApp (e, el) | e = SELF; ".("; qid = Prim.qualid; ")" -> - Loc.tag ~loc:!@loc @@ CTacPrj (e, RelId qid) + CAst.make ~loc:!@loc @@ CTacPrj (e, RelId qid) | e = SELF; ".("; qid = Prim.qualid; ")"; ":="; r = tac2expr LEVEL "5" -> - Loc.tag ~loc:!@loc @@ CTacSet (e, RelId qid, r) ] + CAst.make ~loc:!@loc @@ CTacSet (e, RelId qid, r) ] | "0" [ "("; a = SELF; ")" -> a | "("; a = SELF; ":"; t = tac2type; ")" -> - Loc.tag ~loc:!@loc @@ CTacCnv (a, t) + CAst.make ~loc:!@loc @@ CTacCnv (a, t) | "()" -> - Loc.tag ~loc:!@loc @@ CTacCst (AbsKn (Tuple 0)) + CAst.make ~loc:!@loc @@ CTacCst (AbsKn (Tuple 0)) | "("; ")" -> - Loc.tag ~loc:!@loc @@ CTacCst (AbsKn (Tuple 0)) + CAst.make ~loc:!@loc @@ CTacCst (AbsKn (Tuple 0)) | "["; a = LIST0 tac2expr LEVEL "5" SEP ";"; "]" -> Tac2quote.of_list ~loc:!@loc (fun x -> x) a | "{"; a = tac2rec_fieldexprs; "}" -> - Loc.tag ~loc:!@loc @@ CTacRec a + CAst.make ~loc:!@loc @@ CTacRec a | a = tactic_atom -> a ] ] ; @@ -204,14 +204,14 @@ GEXTEND Gram [ [ "'"; id = Prim.ident -> id ] ] ; tactic_atom: - [ [ n = Prim.integer -> Loc.tag ~loc:!@loc @@ CTacAtm (AtmInt n) - | s = Prim.string -> Loc.tag ~loc:!@loc @@ CTacAtm (AtmStr s) + [ [ n = Prim.integer -> CAst.make ~loc:!@loc @@ CTacAtm (AtmInt n) + | s = Prim.string -> CAst.make ~loc:!@loc @@ CTacAtm (AtmStr s) | id = Prim.qualid -> - if Tac2env.is_constructor (snd id) then - Loc.tag ~loc:!@loc @@ CTacCst (RelId id) + if Tac2env.is_constructor id.CAst.v then + CAst.make ~loc:!@loc @@ CTacCst (RelId id) else - Loc.tag ~loc:!@loc @@ CTacRef (RelId id) - | "@"; id = Prim.ident -> Tac2quote.of_ident (Loc.tag ~loc:!@loc id) + CAst.make ~loc:!@loc @@ CTacRef (RelId id) + | "@"; id = Prim.ident -> Tac2quote.of_ident (CAst.make ~loc:!@loc id) | "&"; id = lident -> Tac2quote.of_hyp ~loc:!@loc id | "'"; c = Constr.constr -> inj_open_constr !@loc c | IDENT "constr"; ":"; "("; c = Constr.lconstr; ")" -> Tac2quote.of_constr c @@ -227,7 +227,7 @@ GEXTEND Gram let (pat, fn) = binder in let te = match fn with | None -> te - | Some args -> Loc.tag ~loc:!@loc @@ CTacFun (args, te) + | Some args -> CAst.make ~loc:!@loc @@ CTacFun (args, te) in (pat, te) ] ] @@ -235,42 +235,42 @@ GEXTEND Gram let_binder: [ [ pats = LIST1 input_fun -> match pats with - | [(_, CPatVar _) as pat] -> (pat, None) - | ((_, CPatVar (Name id)) as pat) :: args -> (pat, Some args) + | [{CAst.v=CPatVar _} as pat] -> (pat, None) + | ({CAst.v=CPatVar (Name id)} as pat) :: args -> (pat, Some args) | [pat] -> (pat, None) | _ -> CErrors.user_err ~loc:!@loc (str "Invalid pattern") ] ] ; tac2type: [ "5" RIGHTA - [ t1 = tac2type; "->"; t2 = tac2type -> Loc.tag ~loc:!@loc @@ CTypArrow (t1, t2) ] + [ t1 = tac2type; "->"; t2 = tac2type -> CAst.make ~loc:!@loc @@ CTypArrow (t1, t2) ] | "2" [ t = tac2type; "*"; tl = LIST1 tac2type LEVEL "1" SEP "*" -> let tl = t :: tl in - Loc.tag ~loc:!@loc @@ CTypRef (AbsKn (Tuple (List.length tl)), tl) ] + CAst.make ~loc:!@loc @@ CTypRef (AbsKn (Tuple (List.length tl)), tl) ] | "1" LEFTA - [ t = SELF; qid = Prim.qualid -> Loc.tag ~loc:!@loc @@ CTypRef (RelId qid, [t]) ] + [ t = SELF; qid = Prim.qualid -> CAst.make ~loc:!@loc @@ CTypRef (RelId qid, [t]) ] | "0" [ "("; t = tac2type LEVEL "5"; ")" -> t - | id = typ_param -> Loc.tag ~loc:!@loc @@ CTypVar (Name id) - | "_" -> Loc.tag ~loc:!@loc @@ CTypVar Anonymous - | qid = Prim.qualid -> Loc.tag ~loc:!@loc @@ CTypRef (RelId qid, []) + | id = typ_param -> CAst.make ~loc:!@loc @@ CTypVar (Name id) + | "_" -> CAst.make ~loc:!@loc @@ CTypVar Anonymous + | qid = Prim.qualid -> CAst.make ~loc:!@loc @@ CTypRef (RelId qid, []) | "("; p = LIST1 tac2type LEVEL "5" SEP ","; ")"; qid = Prim.qualid -> - Loc.tag ~loc:!@loc @@ CTypRef (RelId qid, p) ] + CAst.make ~loc:!@loc @@ CTypRef (RelId qid, p) ] ]; locident: - [ [ id = Prim.ident -> Loc.tag ~loc:!@loc id ] ] + [ [ id = Prim.ident -> CAst.make ~loc:!@loc id ] ] ; binder: - [ [ "_" -> Loc.tag ~loc:!@loc Anonymous - | l = Prim.ident -> Loc.tag ~loc:!@loc (Name l) ] ] + [ [ "_" -> CAst.make ~loc:!@loc Anonymous + | l = Prim.ident -> CAst.make ~loc:!@loc (Name l) ] ] ; input_fun: [ [ b = tac2pat LEVEL "0" -> b ] ] ; tac2def_body: [ [ name = binder; it = LIST0 input_fun; ":="; e = tac2expr -> - let e = if List.is_empty it then e else Loc.tag ~loc:!@loc @@ CTacFun (it, e) in + let e = if List.is_empty it then e else CAst.make ~loc:!@loc @@ CTacFun (it, e) in (name, e) ] ] ; @@ -319,8 +319,8 @@ GEXTEND Gram ; tac2typ_prm: [ [ -> [] - | id = typ_param -> [Loc.tag ~loc:!@loc id] - | "("; ids = LIST1 [ id = typ_param -> Loc.tag ~loc:!@loc id ] SEP "," ;")" -> ids + | id = typ_param -> [CAst.make ~loc:!@loc id] + | "("; ids = LIST1 [ id = typ_param -> CAst.make ~loc:!@loc id ] SEP "," ;")" -> ids ] ] ; tac2typ_def: @@ -345,13 +345,13 @@ GEXTEND Gram ] ] ; syn_node: - [ [ "_" -> Loc.tag ~loc:!@loc None - | id = Prim.ident -> Loc.tag ~loc:!@loc (Some id) + [ [ "_" -> CAst.make ~loc:!@loc None + | id = Prim.ident -> CAst.make ~loc:!@loc (Some id) ] ] ; sexpr: - [ [ s = Prim.string -> SexprStr (Loc.tag ~loc:!@loc s) - | n = Prim.integer -> SexprInt (Loc.tag ~loc:!@loc n) + [ [ s = Prim.string -> SexprStr (CAst.make ~loc:!@loc s) + | n = Prim.integer -> SexprInt (CAst.make ~loc:!@loc n) | id = syn_node -> SexprRec (!@loc, id, []) | id = syn_node; "("; tok = LIST1 sexpr SEP "," ; ")" -> SexprRec (!@loc, id, tok) @@ -369,11 +369,11 @@ GEXTEND Gram ] ] ; lident: - [ [ id = Prim.ident -> Loc.tag ~loc:!@loc id ] ] + [ [ id = Prim.ident -> CAst.make ~loc:!@loc id ] ] ; globref: [ [ "&"; id = Prim.ident -> Libnames.Ident (Loc.tag ~loc:!@loc id) - | qid = Prim.qualid -> Libnames.Qualid qid + | qid = Prim.qualid -> Libnames.Qualid (Loc.tag ~loc:!@loc qid.CAst.v) ] ] ; END @@ -390,38 +390,38 @@ GEXTEND Gram q_destruction_arg q_reference q_with_bindings q_constr_matching q_goal_matching q_hintdb q_move_location q_pose q_assert; anti: - [ [ "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) ] ] + [ [ "$"; id = Prim.ident -> QAnti (CAst.make ~loc:!@loc id) ] ] ; ident_or_anti: [ [ id = lident -> QExpr id - | "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) + | "$"; id = Prim.ident -> QAnti (CAst.make ~loc:!@loc id) ] ] ; lident: - [ [ id = Prim.ident -> Loc.tag ~loc:!@loc id ] ] + [ [ id = Prim.ident -> CAst.make ~loc:!@loc id ] ] ; lnatural: - [ [ n = Prim.natural -> Loc.tag ~loc:!@loc n ] ] + [ [ n = Prim.natural -> CAst.make ~loc:!@loc n ] ] ; q_ident: [ [ id = ident_or_anti -> id ] ] ; qhyp: [ [ x = anti -> x - | n = lnatural -> QExpr (Loc.tag ~loc:!@loc @@ QAnonHyp n) - | id = lident -> QExpr (Loc.tag ~loc:!@loc @@ QNamedHyp id) + | n = lnatural -> QExpr (CAst.make ~loc:!@loc @@ QAnonHyp n) + | id = lident -> QExpr (CAst.make ~loc:!@loc @@ QNamedHyp id) ] ] ; simple_binding: [ [ "("; h = qhyp; ":="; c = Constr.lconstr; ")" -> - Loc.tag ~loc:!@loc (h, c) + CAst.make ~loc:!@loc (h, c) ] ] ; bindings: [ [ test_lpar_idnum_coloneq; bl = LIST1 simple_binding -> - Loc.tag ~loc:!@loc @@ QExplicitBindings bl + CAst.make ~loc:!@loc @@ QExplicitBindings bl | bl = LIST1 Constr.constr -> - Loc.tag ~loc:!@loc @@ QImplicitBindings bl + CAst.make ~loc:!@loc @@ QImplicitBindings bl ] ] ; q_bindings: @@ -431,53 +431,53 @@ GEXTEND Gram [ [ bl = with_bindings -> bl ] ] ; intropatterns: - [ [ l = LIST0 nonsimple_intropattern -> Loc.tag ~loc:!@loc l ]] + [ [ l = LIST0 nonsimple_intropattern -> CAst.make ~loc:!@loc l ]] ; (* ne_intropatterns: *) (* [ [ l = LIST1 nonsimple_intropattern -> l ]] *) (* ; *) or_and_intropattern: - [ [ "["; tc = LIST1 intropatterns SEP "|"; "]" -> Loc.tag ~loc:!@loc @@ QIntroOrPattern tc - | "()" -> Loc.tag ~loc:!@loc @@ QIntroAndPattern (Loc.tag ~loc:!@loc []) - | "("; si = simple_intropattern; ")" -> Loc.tag ~loc:!@loc @@ QIntroAndPattern (Loc.tag ~loc:!@loc [si]) + [ [ "["; tc = LIST1 intropatterns SEP "|"; "]" -> CAst.make ~loc:!@loc @@ QIntroOrPattern tc + | "()" -> CAst.make ~loc:!@loc @@ QIntroAndPattern (CAst.make ~loc:!@loc []) + | "("; si = simple_intropattern; ")" -> CAst.make ~loc:!@loc @@ QIntroAndPattern (CAst.make ~loc:!@loc [si]) | "("; si = simple_intropattern; ","; tc = LIST1 simple_intropattern SEP "," ; ")" -> - Loc.tag ~loc:!@loc @@ QIntroAndPattern (Loc.tag ~loc:!@loc (si::tc)) + CAst.make ~loc:!@loc @@ QIntroAndPattern (CAst.make ~loc:!@loc (si::tc)) | "("; si = simple_intropattern; "&"; tc = LIST1 simple_intropattern SEP "&" ; ")" -> (* (A & B & C) is translated into (A,(B,C)) *) let rec pairify = function - | ([]|[_]|[_;_]) as l -> Loc.tag ~loc:!@loc l + | ([]|[_]|[_;_]) as l -> CAst.make ~loc:!@loc l | t::q -> let q = - Loc.tag ~loc:!@loc @@ - QIntroAction (Loc.tag ~loc:!@loc @@ - QIntroOrAndPattern (Loc.tag ~loc:!@loc @@ + CAst.make ~loc:!@loc @@ + QIntroAction (CAst.make ~loc:!@loc @@ + QIntroOrAndPattern (CAst.make ~loc:!@loc @@ QIntroAndPattern (pairify q))) in - Loc.tag ~loc:!@loc [t; q] - in Loc.tag ~loc:!@loc @@ QIntroAndPattern (pairify (si::tc)) ] ] + CAst.make ~loc:!@loc [t; q] + in CAst.make ~loc:!@loc @@ QIntroAndPattern (pairify (si::tc)) ] ] ; equality_intropattern: - [ [ "->" -> Loc.tag ~loc:!@loc @@ QIntroRewrite true - | "<-" -> Loc.tag ~loc:!@loc @@ QIntroRewrite false - | "[="; tc = intropatterns; "]" -> Loc.tag ~loc:!@loc @@ QIntroInjection tc ] ] + [ [ "->" -> CAst.make ~loc:!@loc @@ QIntroRewrite true + | "<-" -> CAst.make ~loc:!@loc @@ QIntroRewrite false + | "[="; tc = intropatterns; "]" -> CAst.make ~loc:!@loc @@ QIntroInjection tc ] ] ; naming_intropattern: [ [ LEFTQMARK; id = lident -> - Loc.tag ~loc:!@loc @@ QIntroFresh (QExpr id) + CAst.make ~loc:!@loc @@ QIntroFresh (QExpr id) | "?$"; id = lident -> - Loc.tag ~loc:!@loc @@ QIntroFresh (QAnti id) + CAst.make ~loc:!@loc @@ QIntroFresh (QAnti id) | "?" -> - Loc.tag ~loc:!@loc @@ QIntroAnonymous + CAst.make ~loc:!@loc @@ QIntroAnonymous | id = ident_or_anti -> - Loc.tag ~loc:!@loc @@ QIntroIdentifier id + CAst.make ~loc:!@loc @@ QIntroIdentifier id ] ] ; nonsimple_intropattern: [ [ l = simple_intropattern -> l - | "*" -> Loc.tag ~loc:!@loc @@ QIntroForthcoming true - | "**" -> Loc.tag ~loc:!@loc @@ QIntroForthcoming false ]] + | "*" -> CAst.make ~loc:!@loc @@ QIntroForthcoming true + | "**" -> CAst.make ~loc:!@loc @@ QIntroForthcoming false ]] ; simple_intropattern: [ [ pat = simple_intropattern_closed -> @@ -488,13 +488,13 @@ GEXTEND Gram ; simple_intropattern_closed: [ [ pat = or_and_intropattern -> - Loc.tag ~loc:!@loc @@ QIntroAction (Loc.tag ~loc:!@loc @@ QIntroOrAndPattern pat) + CAst.make ~loc:!@loc @@ QIntroAction (CAst.make ~loc:!@loc @@ QIntroOrAndPattern pat) | pat = equality_intropattern -> - Loc.tag ~loc:!@loc @@ QIntroAction pat + CAst.make ~loc:!@loc @@ QIntroAction pat | "_" -> - Loc.tag ~loc:!@loc @@ QIntroAction (Loc.tag ~loc:!@loc @@ QIntroWildcard) + CAst.make ~loc:!@loc @@ QIntroAction (CAst.make ~loc:!@loc @@ QIntroWildcard) | pat = naming_intropattern -> - Loc.tag ~loc:!@loc @@ QIntroNaming pat + CAst.make ~loc:!@loc @@ QIntroNaming pat ] ] ; q_intropatterns: @@ -505,7 +505,7 @@ GEXTEND Gram ; nat_or_anti: [ [ n = lnatural -> QExpr n - | "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) + | "$"; id = Prim.ident -> QAnti (CAst.make ~loc:!@loc id) ] ] ; eqn_ipat: @@ -514,15 +514,15 @@ GEXTEND Gram ] ] ; with_bindings: - [ [ "with"; bl = bindings -> bl | -> Loc.tag ~loc:!@loc @@ QNoBindings ] ] + [ [ "with"; bl = bindings -> bl | -> CAst.make ~loc:!@loc @@ QNoBindings ] ] ; constr_with_bindings: - [ [ c = Constr.constr; l = with_bindings -> Loc.tag ~loc:!@loc @@ (c, l) ] ] + [ [ c = Constr.constr; l = with_bindings -> CAst.make ~loc:!@loc @@ (c, l) ] ] ; destruction_arg: - [ [ n = lnatural -> Loc.tag ~loc:!@loc @@ QElimOnAnonHyp n - | id = lident -> Loc.tag ~loc:!@loc @@ QElimOnIdent id - | c = constr_with_bindings -> Loc.tag ~loc:!@loc @@ QElimOnConstr c + [ [ n = lnatural -> CAst.make ~loc:!@loc @@ QElimOnAnonHyp n + | id = lident -> CAst.make ~loc:!@loc @@ QElimOnIdent id + | c = constr_with_bindings -> CAst.make ~loc:!@loc @@ QElimOnConstr c ] ] ; q_destruction_arg: @@ -534,13 +534,13 @@ GEXTEND Gram ] ] ; occs_nums: - [ [ nl = LIST1 nat_or_anti -> Loc.tag ~loc:!@loc @@ QOnlyOccurrences nl + [ [ nl = LIST1 nat_or_anti -> CAst.make ~loc:!@loc @@ QOnlyOccurrences nl | "-"; n = nat_or_anti; nl = LIST0 nat_or_anti -> - Loc.tag ~loc:!@loc @@ QAllOccurrencesBut (n::nl) + CAst.make ~loc:!@loc @@ QAllOccurrencesBut (n::nl) ] ] ; occs: - [ [ "at"; occs = occs_nums -> occs | -> Loc.tag ~loc:!@loc QAllOccurrences ] ] + [ [ "at"; occs = occs_nums -> occs | -> CAst.make ~loc:!@loc QAllOccurrences ] ] ; hypident: [ [ id = ident_or_anti -> @@ -562,13 +562,13 @@ GEXTEND Gram | hl = LIST0 hypident_occ SEP ","; "|-"; occs = concl_occ -> { q_onhyps = Some hl; q_concl_occs = occs } | hl = LIST0 hypident_occ SEP "," -> - { q_onhyps = Some hl; q_concl_occs = Loc.tag ~loc:!@loc QNoOccurrences } + { q_onhyps = Some hl; q_concl_occs = CAst.make ~loc:!@loc QNoOccurrences } ] ] ; clause: - [ [ "in"; cl = in_clause -> Loc.tag ~loc:!@loc @@ cl + [ [ "in"; cl = in_clause -> CAst.make ~loc:!@loc @@ cl | "at"; occs = occs_nums -> - Loc.tag ~loc:!@loc @@ { q_onhyps = Some []; q_concl_occs = occs } + CAst.make ~loc:!@loc @@ { q_onhyps = Some []; q_concl_occs = occs } ] ] ; q_clause: @@ -576,13 +576,13 @@ GEXTEND Gram ; concl_occ: [ [ "*"; occs = occs -> occs - | -> Loc.tag ~loc:!@loc QNoOccurrences + | -> CAst.make ~loc:!@loc QNoOccurrences ] ] ; induction_clause: [ [ c = destruction_arg; pat = as_or_and_ipat; eq = eqn_ipat; cl = OPT clause -> - Loc.tag ~loc:!@loc @@ { + CAst.make ~loc:!@loc @@ { indcl_arg = c; indcl_eqn = eq; indcl_as = pat; @@ -595,38 +595,38 @@ GEXTEND Gram ; conversion: [ [ c = Constr.constr -> - Loc.tag ~loc:!@loc @@ QConvert c + CAst.make ~loc:!@loc @@ QConvert c | c1 = Constr.constr; "with"; c2 = Constr.constr -> - Loc.tag ~loc:!@loc @@ QConvertWith (c1, c2) + CAst.make ~loc:!@loc @@ QConvertWith (c1, c2) ] ] ; q_conversion: [ [ c = conversion -> c ] ] ; orient: - [ [ "->" -> Loc.tag ~loc:!@loc (Some true) - | "<-" -> Loc.tag ~loc:!@loc (Some false) - | -> Loc.tag ~loc:!@loc None + [ [ "->" -> CAst.make ~loc:!@loc (Some true) + | "<-" -> CAst.make ~loc:!@loc (Some false) + | -> CAst.make ~loc:!@loc None ]] ; rewriter: [ [ "!"; c = constr_with_bindings -> - (Loc.tag ~loc:!@loc @@ QRepeatPlus,c) + (CAst.make ~loc:!@loc @@ QRepeatPlus,c) | ["?"| LEFTQMARK]; c = constr_with_bindings -> - (Loc.tag ~loc:!@loc @@ QRepeatStar,c) + (CAst.make ~loc:!@loc @@ QRepeatStar,c) | n = lnatural; "!"; c = constr_with_bindings -> - (Loc.tag ~loc:!@loc @@ QPrecisely n,c) + (CAst.make ~loc:!@loc @@ QPrecisely n,c) | n = lnatural; ["?" | LEFTQMARK]; c = constr_with_bindings -> - (Loc.tag ~loc:!@loc @@ QUpTo n,c) + (CAst.make ~loc:!@loc @@ QUpTo n,c) | n = lnatural; c = constr_with_bindings -> - (Loc.tag ~loc:!@loc @@ QPrecisely n,c) + (CAst.make ~loc:!@loc @@ QPrecisely n,c) | c = constr_with_bindings -> - (Loc.tag ~loc:!@loc @@ QPrecisely (Loc.tag 1), c) + (CAst.make ~loc:!@loc @@ QPrecisely (CAst.make 1), c) ] ] ; oriented_rewriter: [ [ b = orient; (m, c) = rewriter -> - Loc.tag ~loc:!@loc @@ { + CAst.make ~loc:!@loc @@ { rew_orient = b; rew_repeat = m; rew_equatn = c; @@ -651,52 +651,52 @@ GEXTEND Gram ] ] ; q_dispatch: - [ [ d = tactic_then_gen -> Loc.tag ~loc:!@loc d ] ] + [ [ d = tactic_then_gen -> CAst.make ~loc:!@loc d ] ] ; q_occurrences: [ [ occs = occs -> occs ] ] ; red_flag: - [ [ IDENT "beta" -> Loc.tag ~loc:!@loc @@ QBeta - | IDENT "iota" -> Loc.tag ~loc:!@loc @@ QIota - | IDENT "match" -> Loc.tag ~loc:!@loc @@ QMatch - | IDENT "fix" -> Loc.tag ~loc:!@loc @@ QFix - | IDENT "cofix" -> Loc.tag ~loc:!@loc @@ QCofix - | IDENT "zeta" -> Loc.tag ~loc:!@loc @@ QZeta + [ [ IDENT "beta" -> CAst.make ~loc:!@loc @@ QBeta + | IDENT "iota" -> CAst.make ~loc:!@loc @@ QIota + | IDENT "match" -> CAst.make ~loc:!@loc @@ QMatch + | IDENT "fix" -> CAst.make ~loc:!@loc @@ QFix + | IDENT "cofix" -> CAst.make ~loc:!@loc @@ QCofix + | IDENT "zeta" -> CAst.make ~loc:!@loc @@ QZeta | IDENT "delta"; d = delta_flag -> d ] ] ; refglobal: [ [ "&"; id = Prim.ident -> QExpr (Libnames.Ident (Loc.tag ~loc:!@loc id)) - | qid = Prim.qualid -> QExpr (Libnames.Qualid qid) - | "$"; id = Prim.ident -> QAnti (Loc.tag ~loc:!@loc id) + | qid = Prim.qualid -> QExpr (Libnames.Qualid Loc.(tag ~loc:!@loc qid.CAst.v)) + | "$"; id = Prim.ident -> QAnti (CAst.make ~loc:!@loc id) ] ] ; q_reference: [ [ r = refglobal -> r ] ] ; refglobals: - [ [ gl = LIST1 refglobal -> Loc.tag ~loc:!@loc gl ] ] + [ [ gl = LIST1 refglobal -> CAst.make ~loc:!@loc gl ] ] ; delta_flag: - [ [ "-"; "["; idl = refglobals; "]" -> Loc.tag ~loc:!@loc @@ QDeltaBut idl - | "["; idl = refglobals; "]" -> Loc.tag ~loc:!@loc @@ QConst idl - | -> Loc.tag ~loc:!@loc @@ QDeltaBut (Loc.tag ~loc:!@loc []) + [ [ "-"; "["; idl = refglobals; "]" -> CAst.make ~loc:!@loc @@ QDeltaBut idl + | "["; idl = refglobals; "]" -> CAst.make ~loc:!@loc @@ QConst idl + | -> CAst.make ~loc:!@loc @@ QDeltaBut (CAst.make ~loc:!@loc []) ] ] ; strategy_flag: - [ [ s = LIST1 red_flag -> Loc.tag ~loc:!@loc s + [ [ s = LIST1 red_flag -> CAst.make ~loc:!@loc s | d = delta_flag -> - Loc.tag ~loc:!@loc - [Loc.tag ~loc:!@loc QBeta; Loc.tag ~loc:!@loc QIota; Loc.tag ~loc:!@loc QZeta; d] + CAst.make ~loc:!@loc + [CAst.make ~loc:!@loc QBeta; CAst.make ~loc:!@loc QIota; CAst.make ~loc:!@loc QZeta; d] ] ] ; q_strategy_flag: [ [ flag = strategy_flag -> flag ] ] ; hintdb: - [ [ "*" -> Loc.tag ~loc:!@loc @@ QHintAll - | l = LIST1 ident_or_anti -> Loc.tag ~loc:!@loc @@ QHintDbs l + [ [ "*" -> CAst.make ~loc:!@loc @@ QHintAll + | l = LIST1 ident_or_anti -> CAst.make ~loc:!@loc @@ QHintDbs l ] ] ; q_hintdb: @@ -704,17 +704,17 @@ GEXTEND Gram ; match_pattern: [ [ IDENT "context"; id = OPT Prim.ident; - "["; pat = Constr.lconstr_pattern; "]" -> Loc.tag ~loc:!@loc @@ QConstrMatchContext (id, pat) - | pat = Constr.lconstr_pattern -> Loc.tag ~loc:!@loc @@ QConstrMatchPattern pat ] ] + "["; pat = Constr.lconstr_pattern; "]" -> CAst.make ~loc:!@loc @@ QConstrMatchContext (id, pat) + | pat = Constr.lconstr_pattern -> CAst.make ~loc:!@loc @@ QConstrMatchPattern pat ] ] ; match_rule: [ [ mp = match_pattern; "=>"; tac = tac2expr -> - Loc.tag ~loc:!@loc @@ (mp, tac) + CAst.make ~loc:!@loc @@ (mp, tac) ] ] ; match_list: - [ [ mrl = LIST1 match_rule SEP "|" -> Loc.tag ~loc:!@loc @@ mrl - | "|"; mrl = LIST1 match_rule SEP "|" -> Loc.tag ~loc:!@loc @@ mrl ] ] + [ [ mrl = LIST1 match_rule SEP "|" -> CAst.make ~loc:!@loc @@ mrl + | "|"; mrl = LIST1 match_rule SEP "|" -> CAst.make ~loc:!@loc @@ mrl ] ] ; q_constr_matching: [ [ m = match_list -> m ] ] @@ -724,7 +724,7 @@ GEXTEND Gram ; gmatch_pattern: [ [ "["; hl = LIST0 gmatch_hyp_pattern SEP ","; "|-"; p = match_pattern; "]" -> - Loc.tag ~loc:!@loc @@ { + CAst.make ~loc:!@loc @@ { q_goal_match_concl = p; q_goal_match_hyps = hl; } @@ -732,21 +732,21 @@ GEXTEND Gram ; gmatch_rule: [ [ mp = gmatch_pattern; "=>"; tac = tac2expr -> - Loc.tag ~loc:!@loc @@ (mp, tac) + CAst.make ~loc:!@loc @@ (mp, tac) ] ] ; gmatch_list: - [ [ mrl = LIST1 gmatch_rule SEP "|" -> Loc.tag ~loc:!@loc @@ mrl - | "|"; mrl = LIST1 gmatch_rule SEP "|" -> Loc.tag ~loc:!@loc @@ mrl ] ] + [ [ mrl = LIST1 gmatch_rule SEP "|" -> CAst.make ~loc:!@loc @@ mrl + | "|"; mrl = LIST1 gmatch_rule SEP "|" -> CAst.make ~loc:!@loc @@ mrl ] ] ; q_goal_matching: [ [ m = gmatch_list -> m ] ] ; move_location: - [ [ "at"; IDENT "top" -> Loc.tag ~loc:!@loc @@ QMoveFirst - | "at"; IDENT "bottom" -> Loc.tag ~loc:!@loc @@ QMoveLast - | IDENT "after"; id = ident_or_anti -> Loc.tag ~loc:!@loc @@ QMoveAfter id - | IDENT "before"; id = ident_or_anti -> Loc.tag ~loc:!@loc @@ QMoveBefore id + [ [ "at"; IDENT "top" -> CAst.make ~loc:!@loc @@ QMoveFirst + | "at"; IDENT "bottom" -> CAst.make ~loc:!@loc @@ QMoveLast + | IDENT "after"; id = ident_or_anti -> CAst.make ~loc:!@loc @@ QMoveAfter id + | IDENT "before"; id = ident_or_anti -> CAst.make ~loc:!@loc @@ QMoveBefore id ] ] ; q_move_location: @@ -759,8 +759,8 @@ GEXTEND Gram ; pose: [ [ test_lpar_id_coloneq; "("; id = ident_or_anti; ":="; c = Constr.lconstr; ")" -> - Loc.tag ~loc:!@loc (Some id, c) - | c = Constr.constr; na = as_name -> Loc.tag ~loc:!@loc (na, c) + CAst.make ~loc:!@loc (Some id, c) + | c = Constr.constr; na = as_name -> CAst.make ~loc:!@loc (na, c) ] ] ; q_pose: @@ -778,13 +778,13 @@ GEXTEND Gram ; assertion: [ [ test_lpar_id_coloneq; "("; id = ident_or_anti; ":="; c = Constr.lconstr; ")" -> - Loc.tag ~loc:!@loc (QAssertValue (id, c)) + CAst.make ~loc:!@loc (QAssertValue (id, c)) | test_lpar_id_colon; "("; id = ident_or_anti; ":"; c = Constr.lconstr; ")"; tac = by_tactic -> let loc = !@loc in - let ipat = Loc.tag ~loc @@ QIntroNaming (Loc.tag ~loc @@ QIntroIdentifier id) in - Loc.tag ~loc (QAssertType (Some ipat, c, tac)) + let ipat = CAst.make ~loc @@ QIntroNaming (CAst.make ~loc @@ QIntroIdentifier id) in + CAst.make ~loc (QAssertType (Some ipat, c, tac)) | c = Constr.constr; ipat = as_ipat; tac = by_tactic -> - Loc.tag ~loc:!@loc (QAssertType (ipat, c, tac)) + CAst.make ~loc:!@loc (QAssertType (ipat, c, tac)) ] ] ; q_assert: @@ -801,7 +801,7 @@ GEXTEND Gram let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in CAst.make ~loc:!@loc (CHole (None, IntroAnonymous, Some arg)) | test_ampersand_ident; "&"; id = Prim.ident -> - let tac = Tac2quote.of_exact_hyp ~loc:!@loc (Loc.tag ~loc:!@loc id) in + let tac = Tac2quote.of_exact_hyp ~loc:!@loc (CAst.make ~loc:!@loc id) in let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in CAst.make ~loc:!@loc (CHole (None, IntroAnonymous, Some arg)) | test_dollar_ident; "$"; id = Prim.ident -> diff --git a/src/tac2core.ml b/src/tac2core.ml index c16e72b801..c6c3f26b6b 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -516,7 +516,7 @@ let () = define3 "constr_in_context" ident constr closure begin fun id t c -> Proofview.Unsafe.tclEVARS sigma >>= fun () -> Proofview.Unsafe.tclSETGOALS [Proofview.with_empty_state evk] >>= fun () -> thaw c >>= fun _ -> - Proofview.Unsafe.tclSETGOALS [Proofview.with_empty_state (Proofview.Goal.goal (Proofview.Goal.assume gl))] >>= fun () -> + Proofview.Unsafe.tclSETGOALS [Proofview.with_empty_state (Proofview.Goal.goal gl)] >>= fun () -> let args = List.map (fun d -> EConstr.mkVar (get_id d)) (EConstr.named_context env) in let args = Array.of_list (EConstr.mkRel 1 :: args) in let ans = EConstr.mkEvar (evk, args) in @@ -744,7 +744,6 @@ end let () = define1 "refine" closure begin fun c -> let c = thaw c >>= fun c -> Proofview.tclUNIT ((), Value.to_constr c) in Proofview.Goal.enter begin fun gl -> - let gl = Proofview.Goal.assume gl in Refine.generic_refine ~typecheck:true c gl end >>= fun () -> return v_unit end @@ -1023,10 +1022,10 @@ let () = let add_scope s f = Tac2entries.register_scope (Id.of_string s) f -let rec pr_scope = function -| SexprStr (_, s) -> qstring s -| SexprInt (_, n) -> Pp.int n -| SexprRec (_, (_, na), args) -> +let rec pr_scope = let open CAst in function +| SexprStr {v=s} -> qstring s +| SexprInt {v=n} -> Pp.int n +| SexprRec (_, {v=na}, args) -> let na = match na with | None -> str "_" | Some id -> Id.print id @@ -1037,27 +1036,29 @@ let scope_fail s args = let args = str "(" ++ prlist_with_sep (fun () -> str ", ") pr_scope args ++ str ")" in CErrors.user_err (str "Invalid arguments " ++ args ++ str " in scope " ++ str s) -let q_unit = Loc.tag @@ CTacCst (AbsKn (Tuple 0)) +let q_unit = CAst.make @@ CTacCst (AbsKn (Tuple 0)) let add_generic_scope s entry arg = let parse = function | [] -> let scope = Extend.Aentry entry in - let act x = Loc.tag @@ CTacExt (arg, x) in + let act x = CAst.make @@ CTacExt (arg, x) in Tac2entries.ScopeRule (scope, act) | arg -> scope_fail s arg in add_scope s parse +open CAst + let () = add_scope "keyword" begin function -| [SexprStr (loc, s)] -> +| [SexprStr {loc;v=s}] -> let scope = Extend.Atoken (Tok.KEYWORD s) in Tac2entries.ScopeRule (scope, (fun _ -> q_unit)) | arg -> scope_fail "keyword" arg end let () = add_scope "terminal" begin function -| [SexprStr (loc, s)] -> +| [SexprStr {loc;v=s}] -> let scope = Extend.Atoken (CLexer.terminal s) in Tac2entries.ScopeRule (scope, (fun _ -> q_unit)) | arg -> scope_fail "terminal" arg @@ -1069,7 +1070,7 @@ let () = add_scope "list0" begin function let scope = Extend.Alist0 scope in let act l = Tac2quote.of_list act l in Tac2entries.ScopeRule (scope, act) -| [tok; SexprStr (_, str)] -> +| [tok; SexprStr {v=str}] -> let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in let sep = Extend.Atoken (CLexer.terminal str) in let scope = Extend.Alist0sep (scope, sep) in @@ -1084,7 +1085,7 @@ let () = add_scope "list1" begin function let scope = Extend.Alist1 scope in let act l = Tac2quote.of_list act l in Tac2entries.ScopeRule (scope, act) -| [tok; SexprStr (_, str)] -> +| [tok; SexprStr {v=str}] -> let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in let sep = Extend.Atoken (CLexer.terminal str) in let scope = Extend.Alist1sep (scope, sep) in @@ -1099,9 +1100,9 @@ let () = add_scope "opt" begin function let scope = Extend.Aopt scope in let act opt = match opt with | None -> - Loc.tag @@ CTacCst (AbsKn (Other Core.c_none)) + CAst.make @@ CTacCst (AbsKn (Other Core.c_none)) | Some x -> - Loc.tag @@ CTacApp (Loc.tag @@ CTacCst (AbsKn (Other Core.c_some)), [act x]) + CAst.make @@ CTacApp (CAst.make @@ CTacCst (AbsKn (Other Core.c_some)), [act x]) in Tac2entries.ScopeRule (scope, act) | arg -> scope_fail "opt" arg @@ -1129,7 +1130,7 @@ let () = add_scope "tactic" begin function let scope = Extend.Aentryl (tac2expr, 5) in let act tac = tac in Tac2entries.ScopeRule (scope, act) -| [SexprInt (loc, n)] as arg -> +| [SexprInt {loc;v=n}] as arg -> let () = if n < 0 || n > 6 then scope_fail "tactic" arg in let scope = Extend.Aentryl (tac2expr, n) in let act tac = tac in diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 1631880c71..e4bddf439b 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -8,6 +8,7 @@ open Pp open Util +open CAst open CErrors open Names open Libnames @@ -277,62 +278,61 @@ let fresh_var avoid x = in Namegen.next_ident_away_from (Id.of_string x) bad -let extract_pattern_type (loc, p as pat) = match p with +let extract_pattern_type ({loc;v=p} as pat) = match p with | CPatCnv (pat, ty) -> pat, Some ty | CPatVar _ | CPatRef _ -> pat, None (** Mangle recursive tactics *) let inline_rec_tactic tactics = - let avoid = List.fold_left (fun accu ((_, id), _) -> Id.Set.add id accu) Id.Set.empty tactics in - let map (id, e) = match snd e with + let avoid = List.fold_left (fun accu ({v=id}, _) -> Id.Set.add id accu) Id.Set.empty tactics in + let map (id, e) = match e.v with | CTacFun (pat, _) -> (id, List.map extract_pattern_type pat, e) | _ -> - let loc, _ = id in - user_err ?loc (str "Recursive tactic definitions must be functions") + user_err ?loc:id.loc (str "Recursive tactic definitions must be functions") in let tactics = List.map map tactics in let map (id, pat, e) = let fold_var (avoid, ans) (pat, _) = let id = fresh_var avoid "x" in - let loc = loc_of_patexpr pat in - (Id.Set.add id avoid, Loc.tag ?loc id :: ans) + let loc = pat.loc in + (Id.Set.add id avoid, CAst.make ?loc id :: ans) in (** Fresh variables to abstract over the function patterns *) let _, vars = List.fold_left fold_var (avoid, []) pat in - let map_body ((loc, id), _, e) = (Loc.tag ?loc @@ CPatVar (Name id)), e in + let map_body ({loc;v=id}, _, e) = CAst.(make ?loc @@ CPatVar (Name id)), e in let bnd = List.map map_body tactics in - let pat_of_id (loc, id) = (Loc.tag ?loc @@ CPatVar (Name id)) in - let var_of_id (loc, id) = - let qid = (loc, qualid_of_ident id) in - Loc.tag ?loc @@ CTacRef (RelId qid) + let pat_of_id {loc;v=id} = CAst.make ?loc @@ CPatVar (Name id) in + let var_of_id {loc;v=id} = + let qid = CAst.make ?loc @@ qualid_of_ident id in + CAst.make ?loc @@ CTacRef (RelId qid) in - let loc0 = loc_of_tacexpr e in + let loc0 = e.loc in let vpat = List.map pat_of_id vars in let varg = List.map var_of_id vars in - let e = Loc.tag ?loc:loc0 @@ CTacLet (true, bnd, Loc.tag ?loc:loc0 @@ CTacApp (var_of_id id, varg)) in - (id, Loc.tag ?loc:loc0 @@ CTacFun (vpat, e)) + let e = CAst.make ?loc:loc0 @@ CTacLet (true, bnd, CAst.make ?loc:loc0 @@ CTacApp (var_of_id id, varg)) in + (id, CAst.make ?loc:loc0 @@ CTacFun (vpat, e)) in List.map map tactics -let check_lowercase (loc, id) = +let check_lowercase {loc;v=id} = if Tac2env.is_constructor (Libnames.qualid_of_ident id) then user_err ?loc (str "The identifier " ++ Id.print id ++ str " must be lowercase") let register_ltac ?(local = false) ?(mut = false) isrec tactics = - let map ((loc, na), e) = + let map ({loc;v=na}, e) = let id = match na with | Anonymous -> user_err ?loc (str "Tactic definition must have a name") | Name id -> id in - let () = check_lowercase (loc, id) in - ((loc, id), e) + let () = check_lowercase CAst.(make ?loc id) in + (CAst.(make ?loc id), e) in let tactics = List.map map tactics in let tactics = if isrec then inline_rec_tactic tactics else tactics in - let map ((loc, id), e) = + let map ({loc;v=id}, e) = let (e, t) = intern ~strict:true e in let () = if not (is_value e) then @@ -360,23 +360,23 @@ let register_ltac ?(local = false) ?(mut = false) isrec tactics = in List.iter iter defs -let qualid_to_ident (loc, qid) = +let qualid_to_ident {loc;v=qid} = let (dp, id) = Libnames.repr_qualid qid in - if DirPath.is_empty dp then (loc, id) + if DirPath.is_empty dp then CAst.make ?loc id else user_err ?loc (str "Identifier expected") let register_typedef ?(local = false) isrec types = - let same_name ((_, id1), _) ((_, id2), _) = Id.equal id1 id2 in + let same_name ({v=id1}, _) ({v=id2}, _) = Id.equal id1 id2 in let () = match List.duplicates same_name types with | [] -> () - | ((loc, id), _) :: _ -> + | ({loc;v=id}, _) :: _ -> user_err ?loc (str "Multiple definition of the type name " ++ Id.print id) in - let check ((loc, id), (params, def)) = - let same_name (_, id1) (_, id2) = Id.equal id1 id2 in + let check ({loc;v=id}, (params, def)) = + let same_name {v=id1} {v=id2} = Id.equal id1 id2 in let () = match List.duplicates same_name params with | [] -> () - | (loc, id) :: _ -> + | {loc;v=id} :: _ -> user_err ?loc (str "The type parameter " ++ Id.print id ++ str " occurs several times") in @@ -409,13 +409,13 @@ let register_typedef ?(local = false) isrec types = let () = List.iter check types in let self = if isrec then - let fold accu ((_, id), (params, _)) = + let fold accu ({v=id}, (params, _)) = Id.Map.add id (Lib.make_kn id, List.length params) accu in List.fold_left fold Id.Map.empty types else Id.Map.empty in - let map ((_, id), def) = + let map ({v=id}, def) = let typdef = { typdef_local = local; typdef_expr = intern_typedef self def; @@ -426,7 +426,7 @@ let register_typedef ?(local = false) isrec types = let iter (id, def) = ignore (Lib.add_leaf id (inTypDef def)) in List.iter iter types -let register_primitive ?(local = false) (loc, id) t ml = +let register_primitive ?(local = false) {loc;v=id} t ml = let t = intern_open_type t in let rec count_arrow = function | GTypArrow (_, t) -> 1 + count_arrow t @@ -453,7 +453,7 @@ let register_primitive ?(local = false) (loc, id) t ml = } in ignore (Lib.add_leaf id (inTacDef def)) -let register_open ?(local = false) (loc, qid) (params, def) = +let register_open ?(local = false) {loc;v=qid} (params, def) = let kn = try Tac2env.locate_type qid with Not_found -> @@ -496,14 +496,13 @@ let register_open ?(local = false) (loc, qid) (params, def) = let register_type ?local isrec types = match types with | [qid, true, def] -> - let (loc, _) = qid in + let {loc} = qid in let () = if isrec then user_err ?loc (str "Extensions cannot be recursive") in register_open ?local qid def | _ -> let map (qid, redef, def) = - let (loc, _) = qid in let () = if redef then - user_err ?loc (str "Types can only be extended one by one") + user_err ?loc:qid.loc (str "Types can only be extended one by one") in (qualid_to_ident qid, def) in @@ -530,26 +529,26 @@ module ParseToken = struct let loc_of_token = function -| SexprStr (loc, _) -> loc -| SexprInt (loc, _) -> loc +| SexprStr {loc} -> loc +| SexprInt {loc} -> loc | SexprRec (loc, _, _) -> Some loc let parse_scope = function -| SexprRec (_, (loc, Some id), toks) -> +| SexprRec (_, {loc;v=Some id}, toks) -> if Id.Map.mem id !scope_table then Id.Map.find id !scope_table toks else CErrors.user_err ?loc (str "Unknown scope" ++ spc () ++ Names.Id.print id) -| SexprStr (_, str) -> - let v_unit = Loc.tag @@ CTacCst (AbsKn (Tuple 0)) in +| SexprStr {v=str} -> + let v_unit = CAst.make @@ CTacCst (AbsKn (Tuple 0)) in ScopeRule (Extend.Atoken (Tok.IDENT str), (fun _ -> v_unit)) | tok -> let loc = loc_of_token tok in CErrors.user_err ?loc (str "Invalid parsing token") let parse_token = function -| SexprStr (_, s) -> TacTerm s -| SexprRec (_, (_, na), [tok]) -> +| SexprStr {v=s} -> TacTerm s +| SexprRec (_, {v=na}, [tok]) -> let na = match na with None -> Anonymous | Some id -> Name id in let scope = parse_scope tok in TacNonTerm (na, scope) @@ -591,11 +590,10 @@ let perform_notation syn st = let KRule (rule, act) = get_rule tok in let mk loc args = let map (na, e) = - let loc = loc_of_tacexpr e in - ((Loc.tag ?loc @@ CPatVar na), e) + ((CAst.make ?loc:e.loc @@ CPatVar na), e) in let bnd = List.map map args in - Loc.tag ~loc @@ CTacLet (false, bnd, syn.synext_exp) + CAst.make ~loc @@ CTacLet (false, bnd, syn.synext_exp) in let rule = Extend.Rule (rule, act mk) in let lev = match syn.synext_lev with @@ -659,9 +657,9 @@ let inTac2Abbreviation : abbreviation -> obj = classify_function = classify_abbreviation} let register_notation ?(local = false) tkn lev body = match tkn, lev with -| [SexprRec (_, (loc, Some id), [])], None -> +| [SexprRec (_, {loc;v=Some id}, [])], None -> (** Tactic abbreviation *) - let () = check_lowercase (loc, id) in + let () = check_lowercase CAst.(make ?loc id) in let body = Tac2intern.globalize Id.Set.empty body in let abbr = { abbr_body = body } in ignore (Lib.add_leaf id (inTac2Abbreviation abbr)) @@ -780,7 +778,7 @@ let register_struct ?local str = match str with | StrTyp (isrec, t) -> register_type ?local isrec t | StrPrm (id, t, ml) -> register_primitive ?local id t ml | StrSyn (tok, lev, e) -> register_notation ?local tok lev e -| StrMut (qid, e) -> register_redefinition ?local qid e +| StrMut (qid, e) -> register_redefinition ?local CAst.(qid.loc, qid.v) e | StrRun e -> perform_eval e (** Toplevel exception *) @@ -876,7 +874,7 @@ let solve default tac = if not status then Feedback.feedback Feedback.AddedAxiom let call ~default e = - let loc = loc_of_tacexpr e in + let loc = e.loc in let (e, t) = intern ~strict:false e in let () = check_unit ?loc t in let tac = Tac2interp.interp Tac2interp.empty_environment e in diff --git a/src/tac2entries.mli b/src/tac2entries.mli index a92e149a85..cfb58ea383 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Loc open Names open Libnames open Tac2expr @@ -14,13 +13,13 @@ open Tac2expr (** {5 Toplevel definitions} *) val register_ltac : ?local:bool -> ?mut:bool -> rec_flag -> - (Name.t located * raw_tacexpr) list -> unit + (Misctypes.lname * raw_tacexpr) list -> unit val register_type : ?local:bool -> rec_flag -> - (qualid located * redef_flag * raw_quant_typedef) list -> unit + (qualid CAst.t * redef_flag * raw_quant_typedef) list -> unit val register_primitive : ?local:bool -> - Id.t located -> raw_typexpr -> ml_tactic_name -> unit + Misctypes.lident -> raw_typexpr -> ml_tactic_name -> unit val register_struct : ?local:bool -> strexpr -> unit @@ -63,11 +62,11 @@ val tac2expr : raw_tacexpr Pcoq.Gram.entry open Tac2qexpr -val q_ident : Id.t located or_anti Pcoq.Gram.entry +val q_ident : Id.t CAst.t or_anti Pcoq.Gram.entry val q_bindings : bindings Pcoq.Gram.entry val q_with_bindings : bindings Pcoq.Gram.entry val q_intropattern : intro_pattern Pcoq.Gram.entry -val q_intropatterns : intro_pattern list located Pcoq.Gram.entry +val q_intropatterns : intro_pattern list CAst.t Pcoq.Gram.entry val q_destruction_arg : destruction_arg Pcoq.Gram.entry val q_induction_clause : induction_clause Pcoq.Gram.entry val q_conversion : conversion Pcoq.Gram.entry diff --git a/src/tac2expr.mli b/src/tac2expr.mli index 60f10d360f..ddffd13a31 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -27,7 +27,7 @@ type tacref = | TacAlias of ltac_alias type 'a or_relid = -| RelId of qualid located +| RelId of qualid CAst.t | AbsKn of 'a (** {5 Misc} *) @@ -48,7 +48,7 @@ type raw_typexpr_r = | CTypArrow of raw_typexpr * raw_typexpr | CTypRef of type_constant or_tuple or_relid * raw_typexpr list -and raw_typexpr = raw_typexpr_r located +and raw_typexpr = raw_typexpr_r CAst.t type raw_typedef = | CTydDef of raw_typexpr option @@ -78,7 +78,7 @@ type glb_typedef = type type_scheme = int * int glb_typexpr -type raw_quant_typedef = Id.t located list * raw_typedef +type raw_quant_typedef = Misctypes.lident list * raw_typedef type glb_quant_typedef = int * glb_typedef (** {5 Term syntax} *) @@ -93,7 +93,7 @@ type raw_patexpr_r = | CPatRef of ltac_constructor or_tuple or_relid * raw_patexpr list | CPatCnv of raw_patexpr * raw_typexpr -and raw_patexpr = raw_patexpr_r located +and raw_patexpr = raw_patexpr_r CAst.t type raw_tacexpr_r = | CTacAtm of atom @@ -110,7 +110,7 @@ type raw_tacexpr_r = | CTacSet of raw_tacexpr * ltac_projection or_relid * raw_tacexpr | CTacExt : ('a, _) Tac2dyn.Arg.tag * 'a -> raw_tacexpr_r -and raw_tacexpr = raw_tacexpr_r located +and raw_tacexpr = raw_tacexpr_r CAst.t and raw_taccase = raw_patexpr * raw_tacexpr @@ -152,22 +152,22 @@ type exp_level = | E0 type sexpr = -| SexprStr of string located -| SexprInt of int located -| SexprRec of Loc.t * Id.t option located * sexpr list +| SexprStr of string CAst.t +| SexprInt of int CAst.t +| SexprRec of Loc.t * Id.t option CAst.t * sexpr list (** {5 Toplevel statements} *) type strexpr = -| StrVal of mutable_flag * rec_flag * (Name.t located * raw_tacexpr) list +| StrVal of mutable_flag * rec_flag * (Misctypes.lname * raw_tacexpr) list (** Term definition *) -| StrTyp of rec_flag * (qualid located * redef_flag * raw_quant_typedef) list +| StrTyp of rec_flag * (qualid CAst.t * redef_flag * raw_quant_typedef) list (** Type definition *) -| StrPrm of Id.t located * raw_typexpr * ml_tactic_name +| StrPrm of Misctypes.lident * raw_typexpr * ml_tactic_name (** External definition *) | StrSyn of sexpr list * int option * raw_tacexpr (** Syntactic extensions *) -| StrMut of qualid located * raw_tacexpr +| StrMut of qualid CAst.t * raw_tacexpr (** Redefinition of mutable globals *) | StrRun of raw_tacexpr (** Toplevel evaluation of an expression *) diff --git a/src/tac2intern.ml b/src/tac2intern.ml index dc142043e8..b1dd8f7f51 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -8,6 +8,7 @@ open Pp open Util +open CAst open CErrors open Names open Libnames @@ -172,7 +173,7 @@ let drop_ltac2_env store = let fresh_id env = UF.fresh env.env_cst -let get_alias (loc, id) env = +let get_alias {loc;v=id} env = try Id.Map.find id env.env_als.contents with Not_found -> if env.env_opn then @@ -185,10 +186,6 @@ let push_name id t env = match id with | Anonymous -> env | Name id -> { env with env_var = Id.Map.add id t env.env_var } -let loc_of_tacexpr (loc, _) : Loc.t option = loc - -let loc_of_patexpr (loc, _) : Loc.t option = loc - let error_nargs_mismatch ?loc kn nargs nfound = let cstr = Tac2env.shortest_qualid_of_constructor kn in user_err ?loc (str "Constructor " ++ pr_qualid cstr ++ str " expects " ++ @@ -206,12 +203,12 @@ let rec subst_type subst (t : 'a glb_typexpr) = match t with | GTypRef (qid, args) -> GTypRef (qid, List.map (fun t -> subst_type subst t) args) -let rec intern_type env ((loc, t) : raw_typexpr) : UF.elt glb_typexpr = match t with -| CTypVar (Name id) -> GTypVar (get_alias (Loc.tag ?loc id) env) +let rec intern_type env ({loc;v=t} : raw_typexpr) : UF.elt glb_typexpr = match t with +| CTypVar (Name id) -> GTypVar (get_alias (CAst.make ?loc id) env) | CTypVar Anonymous -> GTypVar (fresh_id env) | CTypRef (rel, args) -> let (kn, nparams) = match rel with - | RelId (loc, qid) -> + | RelId {loc;v=qid} -> let (dp, id) = repr_qualid qid in if DirPath.is_empty dp && Id.Map.mem id env.env_rec then let (kn, n) = Id.Map.find id env.env_rec in @@ -233,9 +230,9 @@ let rec intern_type env ((loc, t) : raw_typexpr) : UF.elt glb_typexpr = match t let nargs = List.length args in let () = if not (Int.equal nparams nargs) then - let loc, qid = match rel with + let {loc;v=qid} = match rel with | RelId lid -> lid - | AbsKn (Other kn) -> loc, shortest_qualid_of_type kn + | AbsKn (Other kn) -> CAst.make ?loc @@ shortest_qualid_of_type kn | AbsKn (Tuple _) -> assert false in user_err ?loc (strbrk "The type constructor " ++ pr_qualid qid ++ @@ -500,10 +497,10 @@ let check_unit ?loc t = let check_redundant_clause = function | [] -> () -| (p, _) :: _ -> warn_redundant_clause ?loc:(loc_of_patexpr p) () +| (p, _) :: _ -> warn_redundant_clause ?loc:p.loc () let get_variable0 mem var = match var with -| RelId (loc, qid) -> +| RelId {loc;v=qid} -> let (dp, id) = repr_qualid qid in if DirPath.is_empty dp && mem id then ArgVar CAst.(make ?loc id) else @@ -520,7 +517,7 @@ let get_variable env var = get_variable0 mem var let get_constructor env var = match var with -| RelId (loc, qid) -> +| RelId {loc;v=qid} -> let c = try Some (Tac2env.locate_constructor qid) with Not_found -> None in begin match c with | Some knc -> Other knc @@ -530,7 +527,7 @@ let get_constructor env var = match var with | AbsKn knc -> knc let get_projection var = match var with -| RelId (loc, qid) -> +| RelId {loc;v=qid} -> let kn = try Tac2env.locate_projection qid with Not_found -> user_err ?loc (pr_qualid qid ++ str " is not a projection") in @@ -556,7 +553,7 @@ type glb_patexpr = | GPatVar of Name.t | GPatRef of ltac_constructor or_tuple * glb_patexpr list -let rec intern_patexpr env (loc, pat) = match pat with +let rec intern_patexpr env {loc;v=pat} = match pat with | CPatVar na -> GPatVar na | CPatRef (qid, pl) -> let kn = get_constructor env qid in @@ -601,7 +598,7 @@ let add_name accu = function | Name id -> Id.Set.add id accu | Anonymous -> accu -let rec ids_of_pattern accu (_, pat) = match pat with +let rec ids_of_pattern accu {v=pat} = match pat with | CPatVar Anonymous -> accu | CPatVar (Name id) -> Id.Set.add id accu | CPatRef (_, pl) -> @@ -609,24 +606,23 @@ let rec ids_of_pattern accu (_, pat) = match pat with | CPatCnv (pat, _) -> ids_of_pattern accu pat let loc_of_relid = function -| RelId (loc, _) -> loc +| RelId {loc} -> loc | AbsKn _ -> None -let extract_pattern_type (loc, p as pat) = match p with +let extract_pattern_type ({loc;v=p} as pat) = match p with | CPatCnv (pat, ty) -> pat, Some ty | CPatVar _ | CPatRef _ -> pat, None (** Expand pattern: [p => t] becomes [x => match x with p => t end] *) let expand_pattern avoid bnd = let fold (avoid, bnd) (pat, t) = - let na, expand = match snd pat with + let na, expand = match pat.v with | CPatVar na -> (** Don't expand variable patterns *) na, None | _ -> - let loc = loc_of_patexpr pat in let id = fresh_var avoid in - let qid = RelId (Loc.tag ?loc (qualid_of_ident id)) in + let qid = RelId (CAst.make ?loc:pat.loc (qualid_of_ident id)) in Name id, Some qid in let avoid = ids_of_pattern avoid pat in @@ -638,7 +634,7 @@ let expand_pattern avoid bnd = | None -> e | Some qid -> let loc = loc_of_relid qid in - Loc.tag ?loc @@ CTacCse (Loc.tag ?loc @@ CTacRef qid, [pat, e]) + CAst.make ?loc @@ CTacCse (CAst.make ?loc @@ CTacRef qid, [pat, e]) in let expand e = List.fold_left fold e bnd in let nas = List.rev_map (fun (na, _, _) -> na) bnd in @@ -648,7 +644,7 @@ let is_alias env qid = match get_variable env qid with | ArgArg (TacAlias _) -> true | ArgVar _ | (ArgArg (TacConstant _)) -> false -let rec intern_rec env (loc, e) = match e with +let rec intern_rec env {loc;v=e} = match e with | CTacAtm atm -> intern_atm env atm | CTacRef qid -> begin match get_variable env qid with @@ -685,10 +681,10 @@ let rec intern_rec env (loc, e) = match e with let (e, t) = intern_rec env (exp e) in let t = List.fold_right (fun t accu -> GTypArrow (t, accu)) tl t in (GTacFun (nas, e), t) -| CTacApp ((loc, CTacCst qid), args) -> +| CTacApp ({loc;v=CTacCst qid}, args) -> let kn = get_constructor env qid in intern_constructor env loc kn args -| CTacApp ((_, CTacRef qid), args) when is_alias env qid -> +| CTacApp ({v=CTacRef qid}, args) when is_alias env qid -> let kn = match get_variable env qid with | ArgArg (TacAlias kn) -> kn | ArgVar _ | (ArgArg (TacConstant _)) -> assert false @@ -696,18 +692,18 @@ let rec intern_rec env (loc, e) = match e with let e = Tac2env.interp_alias kn in let map arg = (** Thunk alias arguments *) - let loc = loc_of_tacexpr arg in - let t_unit = Loc.tag ?loc @@ CTypRef (AbsKn (Tuple 0), []) in - let var = Loc.tag ?loc @@ CPatCnv (Loc.tag ?loc @@ CPatVar Anonymous, t_unit) in - Loc.tag ?loc @@ CTacFun ([var], arg) + let loc = arg.loc in + let t_unit = CAst.make ?loc @@ CTypRef (AbsKn (Tuple 0), []) in + let var = CAst.make ?loc @@ CPatCnv (CAst.make ?loc @@ CPatVar Anonymous, t_unit) in + CAst.make ?loc @@ CTacFun ([var], arg) in let args = List.map map args in - intern_rec env (Loc.tag ?loc @@ CTacApp (e, args)) + intern_rec env (CAst.make ?loc @@ CTacApp (e, args)) | CTacApp (f, args) -> - let loc = loc_of_tacexpr f in + let loc = f.loc in let (f, ft) = intern_rec env f in let fold arg (args, t) = - let loc = loc_of_tacexpr arg in + let loc = arg.loc in let (arg, argt) = intern_rec env arg in (arg :: args, (loc, argt) :: t) in @@ -726,8 +722,7 @@ let rec intern_rec env (loc, e) = match e with if Id.Set.is_empty common then Id.Set.union ids accu else let id = Id.Set.choose common in - let loc = loc_of_patexpr pat in - user_err ?loc (str "Variable " ++ Id.print id ++ str " is bound several \ + user_err ?loc:pat.loc (str "Variable " ++ Id.print id ++ str " is bound several \ times in this matching") in let ids = List.fold_left fold Id.Set.empty el in @@ -739,7 +734,7 @@ let rec intern_rec env (loc, e) = match e with let () = unify ?loc env t tc in (e, tc) | CTacSeq (e1, e2) -> - let loc1 = loc_of_tacexpr e1 in + let loc1 = e1.loc in let (e1, t1) = intern_rec env e1 in let (e2, t2) = intern_rec env e2 in let () = check_elt_unit loc1 env t1 in @@ -750,7 +745,7 @@ let rec intern_rec env (loc, e) = match e with intern_record env loc fs | CTacPrj (e, proj) -> let pinfo = get_projection proj in - let loc = loc_of_tacexpr e in + let loc = e.loc in let (e, t) = intern_rec env e in let subst = Array.init pinfo.pdata_prms (fun _ -> fresh_id env) in let params = Array.map_to_list (fun i -> GTypVar i) subst in @@ -764,7 +759,7 @@ let rec intern_rec env (loc, e) = match e with let () = if not pinfo.pdata_mutb then let loc = match proj with - | RelId (loc, _) -> loc + | RelId {CAst.loc} -> loc | AbsKn _ -> None in user_err ?loc (str "Field is not mutable") @@ -806,10 +801,9 @@ let rec intern_rec env (loc, e) = match e with (e, tpe) and intern_rec_with_constraint env e exp = - let loc = loc_of_tacexpr e in - let (e, t) = intern_rec env e in - let () = unify ?loc env t exp in - e + let (er, t) = intern_rec env e in + let () = unify ?loc:e.loc env t exp in + er and intern_let env loc ids el e = let avoid = Id.Set.union ids (Id.Map.domain env.env_var) in @@ -837,11 +831,10 @@ and intern_let env loc ids el e = and intern_let_rec env loc ids el e = let map env (pat, t, e) = - let (loc, pat) = pat in - let na = match pat with + let na = match pat.v with | CPatVar na -> na | CPatRef _ | CPatCnv _ -> - user_err ?loc (str "This kind of pattern is forbidden in let-rec bindings") + user_err ?loc:pat.loc (str "This kind of pattern is forbidden in let-rec bindings") in let id = fresh_id env in let env = push_name na (monomorphic (GTypVar id)) env in @@ -849,7 +842,7 @@ and intern_let_rec env loc ids el e = in let (env, el) = List.fold_map map env el in let fold (loc, na, tc, e, id) (el, tl) = - let loc_e = loc_of_tacexpr e in + let loc_e = e.loc in let (e, t) = intern_rec env e in let () = if not (is_rec_rhs e) then @@ -891,7 +884,7 @@ and intern_case env loc e pl = (GTacCse (e', Other kn, [||], [||]), GTypVar r) | PKind_variant kn -> let subst, tc = fresh_reftype env kn in - let () = unify ?loc:(loc_of_tacexpr e) env t tc in + let () = unify ?loc:e.loc env t tc in let (nconst, nnonconst, arities) = match kn with | Tuple 0 -> 1, 0, [0] | Tuple n -> 0, 1, [n] @@ -907,9 +900,9 @@ and intern_case env loc e pl = let rec intern_branch = function | [] -> () | (pat, br) :: rem -> - let tbr = match snd pat with + let tbr = match pat.v with | CPatVar (Name _) -> - let loc = loc_of_patexpr pat in + let loc = pat.loc in todo ?loc () | CPatVar Anonymous -> let () = check_redundant_clause rem in @@ -932,7 +925,7 @@ and intern_case env loc e pl = let _ = List.fold_left fill (0, 0) arities in brT | CPatRef (qid, args) -> - let loc = loc_of_patexpr pat in + let loc = pat.loc in let knc = get_constructor env qid in let kn', index, arity = match knc with | Tuple n -> Tuple n, 0, List.init n (fun i -> GTypVar i) @@ -946,8 +939,8 @@ and intern_case env loc e pl = invalid_pattern ?loc kn kn' in let get_id pat = match pat with - | _, CPatVar na -> na - | loc, _ -> todo ?loc () + | {v=CPatVar na} -> na + | {loc} -> todo ?loc () in let ids = List.map get_id args in let nids = List.length ids in @@ -978,7 +971,7 @@ and intern_case env loc e pl = | CPatCnv _ -> user_err ?loc (str "Pattern not handled yet") in - let () = unify ?loc:(loc_of_tacexpr br) env tbr ret in + let () = unify ?loc:br.loc env tbr ret in intern_branch rem in let () = intern_branch pl in @@ -995,7 +988,7 @@ and intern_case env loc e pl = (ce, ret) | PKind_open kn -> let subst, tc = fresh_reftype env (Other kn) in - let () = unify ?loc:(loc_of_tacexpr e) env t tc in + let () = unify ?loc:e.loc env t tc in let ret = GTypVar (fresh_id env) in let rec intern_branch map = function | [] -> @@ -1014,7 +1007,7 @@ and intern_case env loc e pl = | GPatRef _ -> user_err ?loc (str "TODO: Unhandled match case") (** FIXME *) in - let loc = loc_of_patexpr pat in + let loc = pat.loc in let knc = match knc with | Other knc -> knc | Tuple n -> invalid_pattern ?loc (Other kn) (Tuple n) @@ -1080,7 +1073,7 @@ and intern_constructor env loc kn args = match kn with and intern_record env loc fs = let map (proj, e) = let loc = match proj with - | RelId (loc, _) -> loc + | RelId {CAst.loc} -> loc | AbsKn _ -> None in let proj = get_projection proj in @@ -1213,24 +1206,24 @@ let check_subtype t1 t2 = (** Globalization *) let get_projection0 var = match var with -| RelId (loc, qid) -> +| RelId {CAst.loc;v=qid} -> let kn = try Tac2env.locate_projection qid with Not_found -> user_err ?loc (pr_qualid qid ++ str " is not a projection") in kn | AbsKn kn -> kn -let rec globalize ids (loc, er as e) = match er with +let rec globalize ids ({loc;v=er} as e) = match er with | CTacAtm _ -> e | CTacRef ref -> let mem id = Id.Set.mem id ids in begin match get_variable0 mem ref with | ArgVar _ -> e - | ArgArg kn -> Loc.tag ?loc @@ CTacRef (AbsKn kn) + | ArgArg kn -> CAst.make ?loc @@ CTacRef (AbsKn kn) end | CTacCst qid -> let knc = get_constructor () qid in - Loc.tag ?loc @@ CTacCst (AbsKn knc) + CAst.make ?loc @@ CTacCst (AbsKn knc) | CTacFun (bnd, e) -> let fold (pats, accu) pat = let accu = ids_of_pattern accu pat in @@ -1240,11 +1233,11 @@ let rec globalize ids (loc, er as e) = match er with let bnd, ids = List.fold_left fold ([], ids) bnd in let bnd = List.rev bnd in let e = globalize ids e in - Loc.tag ?loc @@ CTacFun (bnd, e) + CAst.make ?loc @@ CTacFun (bnd, e) | CTacApp (e, el) -> let e = globalize ids e in let el = List.map (fun e -> globalize ids e) el in - Loc.tag ?loc @@ CTacApp (e, el) + CAst.make ?loc @@ CTacApp (e, el) | CTacLet (isrec, bnd, e) -> let fold accu (pat, _) = ids_of_pattern accu pat in let ext = List.fold_left fold Id.Set.empty bnd in @@ -1256,34 +1249,34 @@ let rec globalize ids (loc, er as e) = match er with (qid, globalize ids e) in let bnd = List.map map bnd in - Loc.tag ?loc @@ CTacLet (isrec, bnd, e) + CAst.make ?loc @@ CTacLet (isrec, bnd, e) | CTacCnv (e, t) -> let e = globalize ids e in - Loc.tag ?loc @@ CTacCnv (e, t) + CAst.make ?loc @@ CTacCnv (e, t) | CTacSeq (e1, e2) -> let e1 = globalize ids e1 in let e2 = globalize ids e2 in - Loc.tag ?loc @@ CTacSeq (e1, e2) + CAst.make ?loc @@ CTacSeq (e1, e2) | CTacCse (e, bl) -> let e = globalize ids e in let bl = List.map (fun b -> globalize_case ids b) bl in - Loc.tag ?loc @@ CTacCse (e, bl) + CAst.make ?loc @@ CTacCse (e, bl) | CTacRec r -> let map (p, e) = let p = get_projection0 p in let e = globalize ids e in (AbsKn p, e) in - Loc.tag ?loc @@ CTacRec (List.map map r) + CAst.make ?loc @@ CTacRec (List.map map r) | CTacPrj (e, p) -> let e = globalize ids e in let p = get_projection0 p in - Loc.tag ?loc @@ CTacPrj (e, AbsKn p) + CAst.make ?loc @@ CTacPrj (e, AbsKn p) | CTacSet (e, p, e') -> let e = globalize ids e in let p = get_projection0 p in let e' = globalize ids e' in - Loc.tag ?loc @@ CTacSet (e, AbsKn p, e') + CAst.make ?loc @@ CTacSet (e, AbsKn p, e') | CTacExt (tag, arg) -> let arg = str (Tac2dyn.Arg.repr tag) in CErrors.user_err ?loc (str "Cannot globalize generic arguments of type" ++ spc () ++ arg) @@ -1291,16 +1284,16 @@ let rec globalize ids (loc, er as e) = match er with and globalize_case ids (p, e) = (globalize_pattern ids p, globalize ids e) -and globalize_pattern ids (loc, pr as p) = match pr with +and globalize_pattern ids ({loc;v=pr} as p) = match pr with | CPatVar _ -> p | CPatRef (cst, pl) -> let knc = get_constructor () cst in let cst = AbsKn knc in let pl = List.map (fun p -> globalize_pattern ids p) pl in - Loc.tag ?loc @@ CPatRef (cst, pl) + CAst.make ?loc @@ CPatRef (cst, pl) | CPatCnv (pat, ty) -> let pat = globalize_pattern ids pat in - Loc.tag ?loc @@ CPatCnv (pat, ty) + CAst.make ?loc @@ CPatCnv (pat, ty) (** Kernel substitution *) @@ -1407,16 +1400,16 @@ let subst_or_relid subst ref = match ref with let kn' = subst_or_tuple subst_kn subst kn in if kn' == kn then ref else AbsKn kn' -let rec subst_rawtype subst (loc, tr as t) = match tr with +let rec subst_rawtype subst ({loc;v=tr} as t) = match tr with | CTypVar _ -> t | CTypArrow (t1, t2) -> let t1' = subst_rawtype subst t1 in let t2' = subst_rawtype subst t2 in - if t1' == t1 && t2' == t2 then t else Loc.tag ?loc @@ CTypArrow (t1', t2') + if t1' == t1 && t2' == t2 then t else CAst.make ?loc @@ CTypArrow (t1', t2') | CTypRef (ref, tl) -> let ref' = subst_or_relid subst ref in let tl' = List.smartmap (fun t -> subst_rawtype subst t) tl in - if ref' == ref && tl' == tl then t else Loc.tag ?loc @@ CTypRef (ref', tl') + if ref' == ref && tl' == tl then t else CAst.make ?loc @@ CTypRef (ref', tl') let subst_tacref subst ref = match ref with | RelId _ -> ref @@ -1433,35 +1426,35 @@ let subst_projection subst prj = match prj with let kn' = subst_kn subst kn in if kn' == kn then prj else AbsKn kn' -let rec subst_rawpattern subst (loc, pr as p) = match pr with +let rec subst_rawpattern subst ({loc;v=pr} as p) = match pr with | CPatVar _ -> p | CPatRef (c, pl) -> let pl' = List.smartmap (fun p -> subst_rawpattern subst p) pl in let c' = subst_or_relid subst c in - if pl' == pl && c' == c then p else Loc.tag ?loc @@ CPatRef (c', pl') + if pl' == pl && c' == c then p else CAst.make ?loc @@ CPatRef (c', pl') | CPatCnv (pat, ty) -> let pat' = subst_rawpattern subst pat in let ty' = subst_rawtype subst ty in - if pat' == pat && ty' == ty then p else Loc.tag ?loc @@ CPatCnv (pat', ty') + if pat' == pat && ty' == ty then p else CAst.make ?loc @@ CPatCnv (pat', ty') (** Used for notations *) -let rec subst_rawexpr subst (loc, tr as t) = match tr with +let rec subst_rawexpr subst ({loc;v=tr} as t) = match tr with | CTacAtm _ -> t | CTacRef ref -> let ref' = subst_tacref subst ref in - if ref' == ref then t else Loc.tag ?loc @@ CTacRef ref' + if ref' == ref then t else CAst.make ?loc @@ CTacRef ref' | CTacCst ref -> let ref' = subst_or_relid subst ref in - if ref' == ref then t else Loc.tag ?loc @@ CTacCst ref' + if ref' == ref then t else CAst.make ?loc @@ CTacCst ref' | CTacFun (bnd, e) -> let map pat = subst_rawpattern subst pat in let bnd' = List.smartmap map bnd in let e' = subst_rawexpr subst e in - if bnd' == bnd && e' == e then t else Loc.tag ?loc @@ CTacFun (bnd', e') + if bnd' == bnd && e' == e then t else CAst.make ?loc @@ CTacFun (bnd', e') | CTacApp (e, el) -> let e' = subst_rawexpr subst e in let el' = List.smartmap (fun e -> subst_rawexpr subst e) el in - if e' == e && el' == el then t else Loc.tag ?loc @@ CTacApp (e', el') + if e' == e && el' == el then t else CAst.make ?loc @@ CTacApp (e', el') | CTacLet (isrec, bnd, e) -> let map (na, e as p) = let na' = subst_rawpattern subst na in @@ -1470,15 +1463,15 @@ let rec subst_rawexpr subst (loc, tr as t) = match tr with in let bnd' = List.smartmap map bnd in let e' = subst_rawexpr subst e in - if bnd' == bnd && e' == e then t else Loc.tag ?loc @@ CTacLet (isrec, bnd', e') + if bnd' == bnd && e' == e then t else CAst.make ?loc @@ CTacLet (isrec, bnd', e') | CTacCnv (e, c) -> let e' = subst_rawexpr subst e in let c' = subst_rawtype subst c in - if c' == c && e' == e then t else Loc.tag ?loc @@ CTacCnv (e', c') + if c' == c && e' == e then t else CAst.make ?loc @@ CTacCnv (e', c') | CTacSeq (e1, e2) -> let e1' = subst_rawexpr subst e1 in let e2' = subst_rawexpr subst e2 in - if e1' == e1 && e2' == e2 then t else Loc.tag ?loc @@ CTacSeq (e1', e2') + if e1' == e1 && e2' == e2 then t else CAst.make ?loc @@ CTacSeq (e1', e2') | CTacCse (e, bl) -> let map (p, e as x) = let p' = subst_rawpattern subst p in @@ -1487,7 +1480,7 @@ let rec subst_rawexpr subst (loc, tr as t) = match tr with in let e' = subst_rawexpr subst e in let bl' = List.smartmap map bl in - if e' == e && bl' == bl then t else Loc.tag ?loc @@ CTacCse (e', bl') + if e' == e && bl' == bl then t else CAst.make ?loc @@ CTacCse (e', bl') | CTacRec el -> let map (prj, e as p) = let prj' = subst_projection subst prj in @@ -1495,16 +1488,16 @@ let rec subst_rawexpr subst (loc, tr as t) = match tr with if prj' == prj && e' == e then p else (prj', e') in let el' = List.smartmap map el in - if el' == el then t else Loc.tag ?loc @@ CTacRec el' + if el' == el then t else CAst.make ?loc @@ CTacRec el' | CTacPrj (e, prj) -> let prj' = subst_projection subst prj in let e' = subst_rawexpr subst e in - if prj' == prj && e' == e then t else Loc.tag ?loc @@ CTacPrj (e', prj') + if prj' == prj && e' == e then t else CAst.make ?loc @@ CTacPrj (e', prj') | CTacSet (e, prj, r) -> let prj' = subst_projection subst prj in let e' = subst_rawexpr subst e in let r' = subst_rawexpr subst r in - if prj' == prj && e' == e && r' == r then t else Loc.tag ?loc @@ CTacSet (e', prj', r') + if prj' == prj && e' == e && r' == r then t else CAst.make ?loc @@ CTacSet (e', prj', r') | CTacExt _ -> assert false (** Should not be generated by globalization *) (** Registering *) @@ -1520,7 +1513,7 @@ let () = else { env with env_str = false } | Some env -> env in - let loc = loc_of_tacexpr tac in + let loc = tac.loc in let (tac, t) = intern_rec env tac in let () = check_elt_unit loc env t in (ist, tac) diff --git a/src/tac2intern.mli b/src/tac2intern.mli index 4b02f91caa..d646b5cda5 100644 --- a/src/tac2intern.mli +++ b/src/tac2intern.mli @@ -10,9 +10,6 @@ open Names open Mod_subst open Tac2expr -val loc_of_tacexpr : raw_tacexpr -> Loc.t option -val loc_of_patexpr : raw_patexpr -> Loc.t option - val intern : strict:bool -> raw_tacexpr -> glb_tacexpr * type_scheme val intern_typedef : (KerName.t * int) Id.Map.t -> raw_quant_typedef -> glb_quant_typedef val intern_open_type : raw_typexpr -> type_scheme diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index 2f6c97f08b..05b9f4141f 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Loc open Names open Tac2expr @@ -15,65 +14,65 @@ open Tac2expr type 'a or_anti = | QExpr of 'a -| QAnti of Id.t located +| QAnti of Id.t CAst.t type quantified_hypothesis = -| QAnonHyp of int located -| QNamedHyp of Id.t located +| QAnonHyp of int CAst.t +| QNamedHyp of Id.t CAst.t type bindings_r = | QImplicitBindings of Constrexpr.constr_expr list -| QExplicitBindings of (quantified_hypothesis located or_anti * Constrexpr.constr_expr) located list +| QExplicitBindings of (quantified_hypothesis CAst.t or_anti * Constrexpr.constr_expr) CAst.t list | QNoBindings -type bindings = bindings_r located +type bindings = bindings_r CAst.t type intro_pattern_r = | QIntroForthcoming of bool | QIntroNaming of intro_pattern_naming | QIntroAction of intro_pattern_action and intro_pattern_naming_r = -| QIntroIdentifier of Id.t located or_anti -| QIntroFresh of Id.t located or_anti +| QIntroIdentifier of Id.t CAst.t or_anti +| QIntroFresh of Id.t CAst.t or_anti | QIntroAnonymous and intro_pattern_action_r = | QIntroWildcard | QIntroOrAndPattern of or_and_intro_pattern -| QIntroInjection of intro_pattern list located +| QIntroInjection of intro_pattern list CAst.t (* | QIntroApplyOn of Empty.t (** Not implemented yet *) *) | QIntroRewrite of bool and or_and_intro_pattern_r = -| QIntroOrPattern of intro_pattern list located list -| QIntroAndPattern of intro_pattern list located +| QIntroOrPattern of intro_pattern list CAst.t list +| QIntroAndPattern of intro_pattern list CAst.t -and intro_pattern = intro_pattern_r located -and intro_pattern_naming = intro_pattern_naming_r located -and intro_pattern_action = intro_pattern_action_r located -and or_and_intro_pattern = or_and_intro_pattern_r located +and intro_pattern = intro_pattern_r CAst.t +and intro_pattern_naming = intro_pattern_naming_r CAst.t +and intro_pattern_action = intro_pattern_action_r CAst.t +and or_and_intro_pattern = or_and_intro_pattern_r CAst.t type occurrences_r = | QAllOccurrences -| QAllOccurrencesBut of int located or_anti list +| QAllOccurrencesBut of int CAst.t or_anti list | QNoOccurrences -| QOnlyOccurrences of int located or_anti list +| QOnlyOccurrences of int CAst.t or_anti list -type occurrences = occurrences_r located +type occurrences = occurrences_r CAst.t -type hyp_location = (occurrences * Id.t located or_anti) * Locus.hyp_location_flag +type hyp_location = (occurrences * Id.t CAst.t or_anti) * Locus.hyp_location_flag type clause_r = { q_onhyps : hyp_location list option; q_concl_occs : occurrences; } -type clause = clause_r located +type clause = clause_r CAst.t -type constr_with_bindings = (Constrexpr.constr_expr * bindings) located +type constr_with_bindings = (Constrexpr.constr_expr * bindings) CAst.t type destruction_arg_r = | QElimOnConstr of constr_with_bindings -| QElimOnIdent of Id.t located -| QElimOnAnonHyp of int located +| QElimOnIdent of Id.t CAst.t +| QElimOnAnonHyp of int CAst.t -type destruction_arg = destruction_arg_r located +type destruction_arg = destruction_arg_r CAst.t type induction_clause_r = { indcl_arg : destruction_arg; @@ -82,33 +81,33 @@ type induction_clause_r = { indcl_in : clause option; } -type induction_clause = induction_clause_r located +type induction_clause = induction_clause_r CAst.t type conversion_r = | QConvert of Constrexpr.constr_expr | QConvertWith of Constrexpr.constr_expr * Constrexpr.constr_expr -type conversion = conversion_r located +type conversion = conversion_r CAst.t type multi_r = -| QPrecisely of int located -| QUpTo of int located +| QPrecisely of int CAst.t +| QUpTo of int CAst.t | QRepeatStar | QRepeatPlus -type multi = multi_r located +type multi = multi_r CAst.t type rewriting_r = { - rew_orient : bool option located; + rew_orient : bool option CAst.t; rew_repeat : multi; rew_equatn : constr_with_bindings; } -type rewriting = rewriting_r located +type rewriting = rewriting_r CAst.t type dispatch_r = raw_tacexpr option list * (raw_tacexpr option * raw_tacexpr option list) option -type dispatch = dispatch_r located +type dispatch = dispatch_r CAst.t type red_flag_r = | QBeta @@ -117,52 +116,52 @@ type red_flag_r = | QFix | QCofix | QZeta -| QConst of Libnames.reference or_anti list located -| QDeltaBut of Libnames.reference or_anti list located +| QConst of Libnames.reference or_anti list CAst.t +| QDeltaBut of Libnames.reference or_anti list CAst.t -type red_flag = red_flag_r located +type red_flag = red_flag_r CAst.t -type strategy_flag = red_flag list located +type strategy_flag = red_flag list CAst.t type constr_match_pattern_r = | QConstrMatchPattern of Constrexpr.constr_expr | QConstrMatchContext of Id.t option * Constrexpr.constr_expr -type constr_match_pattern = constr_match_pattern_r located +type constr_match_pattern = constr_match_pattern_r CAst.t -type constr_match_branch = (constr_match_pattern * raw_tacexpr) located +type constr_match_branch = (constr_match_pattern * raw_tacexpr) CAst.t -type constr_matching = constr_match_branch list located +type constr_matching = constr_match_branch list CAst.t type goal_match_pattern_r = { q_goal_match_concl : constr_match_pattern; q_goal_match_hyps : (Misctypes.lname * constr_match_pattern) list; } -type goal_match_pattern = goal_match_pattern_r located +type goal_match_pattern = goal_match_pattern_r CAst.t -type goal_match_branch = (goal_match_pattern * raw_tacexpr) located +type goal_match_branch = (goal_match_pattern * raw_tacexpr) CAst.t -type goal_matching = goal_match_branch list located +type goal_matching = goal_match_branch list CAst.t type hintdb_r = | QHintAll -| QHintDbs of Id.t located or_anti list +| QHintDbs of Id.t CAst.t or_anti list -type hintdb = hintdb_r located +type hintdb = hintdb_r CAst.t type move_location_r = -| QMoveAfter of Id.t located or_anti -| QMoveBefore of Id.t located or_anti +| QMoveAfter of Id.t CAst.t or_anti +| QMoveBefore of Id.t CAst.t or_anti | QMoveFirst | QMoveLast -type move_location = move_location_r located +type move_location = move_location_r CAst.t -type pose = (Id.t located or_anti option * Constrexpr.constr_expr) located +type pose = (Id.t CAst.t or_anti option * Constrexpr.constr_expr) CAst.t type assertion_r = | QAssertType of intro_pattern option * Constrexpr.constr_expr * raw_tacexpr option -| QAssertValue of Id.t located or_anti * Constrexpr.constr_expr +| QAssertValue of Id.t CAst.t or_anti * Constrexpr.constr_expr -type assertion = assertion_r located +type assertion = assertion_r CAst.t diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 829f13344c..e986bfc393 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -9,6 +9,7 @@ open Pp open Names open Util +open CAst open Tac2dyn open Tac2expr open Tac2qexpr @@ -38,12 +39,12 @@ let control_core n = kername control_prefix n let pattern_core n = kername pattern_prefix n let global_ref ?loc kn = - Loc.tag ?loc @@ CTacRef (AbsKn (TacConstant kn)) + CAst.make ?loc @@ CTacRef (AbsKn (TacConstant kn)) let constructor ?loc kn args = - let cst = Loc.tag ?loc @@ CTacCst (AbsKn (Other kn)) in + let cst = CAst.make ?loc @@ CTacCst (AbsKn (Other kn)) in if List.is_empty args then cst - else Loc.tag ?loc @@ CTacApp (cst, args) + else CAst.make ?loc @@ CTacApp (cst, args) let std_constructor ?loc name args = constructor ?loc (std_core name) args @@ -53,44 +54,44 @@ let std_proj ?loc name = let thunk e = let t_unit = coq_core "unit" in - let loc = Tac2intern.loc_of_tacexpr e in - let ty = Loc.tag ?loc @@ CTypRef (AbsKn (Other t_unit), []) in - let pat = Loc.tag ?loc @@ CPatVar (Anonymous) in - let pat = Loc.tag ?loc @@ CPatCnv (pat, ty) in - Loc.tag ?loc @@ CTacFun ([pat], e) + let loc = e.loc in + let ty = CAst.make?loc @@ CTypRef (AbsKn (Other t_unit), []) in + let pat = CAst.make ?loc @@ CPatVar (Anonymous) in + let pat = CAst.make ?loc @@ CPatCnv (pat, ty) in + CAst.make ?loc @@ CTacFun ([pat], e) -let of_pair f g (loc, (e1, e2)) = - Loc.tag ?loc @@ CTacApp (Loc.tag ?loc @@ CTacCst (AbsKn (Tuple 2)), [f e1; g e2]) +let of_pair f g {loc;v=(e1, e2)} = + CAst.make ?loc @@ CTacApp (CAst.make ?loc @@ CTacCst (AbsKn (Tuple 2)), [f e1; g e2]) let of_tuple ?loc el = match el with | [] -> - Loc.tag ?loc @@ CTacCst (AbsKn (Tuple 0)) + CAst.make ?loc @@ CTacCst (AbsKn (Tuple 0)) | [e] -> e | el -> let len = List.length el in - Loc.tag ?loc @@ CTacApp (Loc.tag ?loc @@ CTacCst (AbsKn (Tuple len)), el) + CAst.make ?loc @@ CTacApp (CAst.make ?loc @@ CTacCst (AbsKn (Tuple len)), el) -let of_int (loc, n) = - Loc.tag ?loc @@ CTacAtm (AtmInt n) +let of_int {loc;v=n} = + CAst.make ?loc @@ CTacAtm (AtmInt n) let of_option ?loc f opt = match opt with | None -> constructor ?loc (coq_core "None") [] | Some e -> constructor ?loc (coq_core "Some") [f e] let inj_wit ?loc wit x = - Loc.tag ?loc @@ CTacExt (wit, x) + CAst.make ?loc @@ CTacExt (wit, x) -let of_variable (loc, id) = +let of_variable {loc;v=id} = let qid = Libnames.qualid_of_ident id in if Tac2env.is_constructor qid then CErrors.user_err ?loc (str "Invalid identifier") - else Loc.tag ?loc @@ CTacRef (RelId (Loc.tag ?loc qid)) + else CAst.make ?loc @@ CTacRef (RelId (CAst.make ?loc qid)) let of_anti f = function | QExpr x -> f x | QAnti id -> of_variable id -let of_ident (loc, id) = inj_wit ?loc wit_ident id +let of_ident {loc;v=id} = inj_wit ?loc wit_ident id let of_constr c = let loc = Constrexpr_ops.constr_loc c in @@ -109,11 +110,11 @@ let rec of_list ?loc f = function | e :: l -> constructor ?loc (coq_core "::") [f e; of_list ?loc f l] -let of_qhyp (loc, h) = match h with +let of_qhyp {loc;v=h} = match h with | QAnonHyp n -> std_constructor ?loc "AnonHyp" [of_int n] | QNamedHyp id -> std_constructor ?loc "NamedHyp" [of_ident id] -let of_bindings (loc, b) = match b with +let of_bindings {loc;v=b} = match b with | QNoBindings -> std_constructor ?loc "NoBindings" [] | QImplicitBindings tl -> @@ -124,7 +125,7 @@ let of_bindings (loc, b) = match b with let of_constr_with_bindings c = of_pair of_open_constr of_bindings c -let rec of_intro_pattern (loc, pat) = match pat with +let rec of_intro_pattern {loc;v=pat} = match pat with | QIntroForthcoming b -> std_constructor ?loc "IntroForthcoming" [of_bool b] | QIntroNaming iname -> @@ -132,7 +133,7 @@ let rec of_intro_pattern (loc, pat) = match pat with | QIntroAction iact -> std_constructor ?loc "IntroAction" [of_intro_pattern_action iact] -and of_intro_pattern_naming (loc, pat) = match pat with +and of_intro_pattern_naming {loc;v=pat} = match pat with | QIntroIdentifier id -> std_constructor ?loc "IntroIdentifier" [of_anti of_ident id] | QIntroFresh id -> @@ -140,7 +141,7 @@ and of_intro_pattern_naming (loc, pat) = match pat with | QIntroAnonymous -> std_constructor ?loc "IntroAnonymous" [] -and of_intro_pattern_action (loc, pat) = match pat with +and of_intro_pattern_action {loc;v=pat} = match pat with | QIntroWildcard -> std_constructor ?loc "IntroWildcard" [] | QIntroOrAndPattern pat -> @@ -150,13 +151,13 @@ and of_intro_pattern_action (loc, pat) = match pat with | QIntroRewrite b -> std_constructor ?loc "IntroRewrite" [of_bool ?loc b] -and of_or_and_intro_pattern (loc, pat) = match pat with +and of_or_and_intro_pattern {loc;v=pat} = match pat with | QIntroOrPattern ill -> std_constructor ?loc "IntroOrPattern" [of_list ?loc of_intro_patterns ill] | QIntroAndPattern il -> std_constructor ?loc "IntroAndPattern" [of_intro_patterns il] -and of_intro_patterns (loc, l) = +and of_intro_patterns {loc;v=l} = of_list ?loc of_intro_pattern l let of_hyp_location_flag ?loc = function @@ -164,7 +165,7 @@ let of_hyp_location_flag ?loc = function | Locus.InHypTypeOnly -> std_constructor ?loc "InHypTypeOnly" [] | Locus.InHypValueOnly -> std_constructor ?loc "InHypValueOnly" [] -let of_occurrences (loc, occ) = match occ with +let of_occurrences {loc;v=occ} = match occ with | QAllOccurrences -> std_constructor ?loc "AllOccurrences" [] | QAllOccurrencesBut occs -> let map occ = of_anti of_int occ in @@ -183,27 +184,27 @@ let of_hyp_location ?loc ((occs, id), flag) = of_hyp_location_flag ?loc flag; ] -let of_clause (loc, cl) = +let of_clause {loc;v=cl} = let hyps = of_option ?loc (fun l -> of_list ?loc of_hyp_location l) cl.q_onhyps in let concl = of_occurrences cl.q_concl_occs in - Loc.tag ?loc @@ CTacRec ([ + CAst.make ?loc @@ CTacRec ([ std_proj "on_hyps", hyps; std_proj "on_concl", concl; ]) -let of_destruction_arg (loc, arg) = match arg with +let of_destruction_arg {loc;v=arg} = match arg with | QElimOnConstr c -> let arg = thunk (of_constr_with_bindings c) in std_constructor ?loc "ElimOnConstr" [arg] | QElimOnIdent id -> std_constructor ?loc "ElimOnIdent" [of_ident id] | QElimOnAnonHyp n -> std_constructor ?loc "ElimOnAnonHyp" [of_int n] -let of_induction_clause (loc, cl) = +let of_induction_clause {loc;v=cl} = let arg = of_destruction_arg cl.indcl_arg in let eqn = of_option ?loc of_intro_pattern_naming cl.indcl_eqn in let as_ = of_option ?loc of_or_and_intro_pattern cl.indcl_as in let in_ = of_option ?loc of_clause cl.indcl_in in - Loc.tag ?loc @@ CTacRec ([ + CAst.make ?loc @@ CTacRec ([ std_proj "indcl_arg", arg; std_proj "indcl_eqn", eqn; std_proj "indcl_as", as_; @@ -238,24 +239,24 @@ let abstract_vars loc vars tac = | Anonymous -> (n + 1, accu) | Name _ -> let get = global_ref ?loc (kername array_prefix "get") in - let args = [of_variable (loc, id0); of_int (loc, n)] in - let e = Loc.tag ?loc @@ CTacApp (get, args) in - let accu = (Loc.tag ?loc @@ CPatVar na, e) :: accu in + let args = [of_variable CAst.(make ?loc id0); of_int CAst.(make ?loc n)] in + let e = CAst.make ?loc @@ CTacApp (get, args) in + let accu = (CAst.make ?loc @@ CPatVar na, e) :: accu in (n + 1, accu) in let (_, bnd) = List.fold_left build_bindings (0, []) vars in - let tac = Loc.tag ?loc @@ CTacLet (false, bnd, tac) in + let tac = CAst.make ?loc @@ CTacLet (false, bnd, tac) in (Name id0, tac) in - Loc.tag ?loc @@ CTacFun ([Loc.tag ?loc @@ CPatVar na], tac) + CAst.make ?loc @@ CTacFun ([CAst.make ?loc @@ CPatVar na], tac) let of_pattern p = inj_wit ?loc:p.CAst.loc wit_pattern p -let of_conversion (loc, c) = match c with +let of_conversion {loc;v=c} = match c with | QConvert c -> let pat = of_option ?loc of_pattern None in - let c = Loc.tag ?loc @@ CTacFun ([Loc.tag ?loc @@ CPatVar Anonymous], of_constr c) in + let c = CAst.make ?loc @@ CTacFun ([CAst.make ?loc @@ CPatVar Anonymous], of_constr c) in of_tuple ?loc [pat; c] | QConvertWith (pat, c) -> let vars = pattern_vars pat in @@ -266,7 +267,7 @@ let of_conversion (loc, c) = match c with let c = abstract_vars loc vars c in of_tuple [pat; c] -let of_repeat (loc, r) = match r with +let of_repeat {loc;v=r} = match r with | QPrecisely n -> std_constructor ?loc "Precisely" [of_int n] | QUpTo n -> std_constructor ?loc "UpTo" [of_int n] | QRepeatStar -> std_constructor ?loc "RepeatStar" [] @@ -276,14 +277,14 @@ let of_orient loc b = if b then std_constructor ?loc "LTR" [] else std_constructor ?loc "RTL" [] -let of_rewriting (loc, rew) = +let of_rewriting {loc;v=rew} = let orient = - let (loc, orient) = rew.rew_orient in + let {loc;v=orient} = rew.rew_orient in of_option ?loc (fun b -> of_orient loc b) orient in let repeat = of_repeat rew.rew_repeat in let equatn = thunk (of_constr_with_bindings rew.rew_equatn) in - Loc.tag ?loc @@ CTacRec ([ + CAst.make ?loc @@ CTacRec ([ std_proj "rew_orient", orient; std_proj "rew_repeat", repeat; std_proj "rew_equatn", equatn; @@ -291,42 +292,42 @@ let of_rewriting (loc, rew) = let of_hyp ?loc id = let hyp = global_ref ?loc (control_core "hyp") in - Loc.tag ?loc @@ CTacApp (hyp, [of_ident id]) + CAst.make ?loc @@ CTacApp (hyp, [of_ident id]) let of_exact_hyp ?loc id = let refine = global_ref ?loc (control_core "refine") in - Loc.tag ?loc @@ CTacApp (refine, [thunk (of_hyp ?loc id)]) + CAst.make ?loc @@ CTacApp (refine, [thunk (of_hyp ?loc id)]) let of_exact_var ?loc id = let refine = global_ref ?loc (control_core "refine") in - Loc.tag ?loc @@ CTacApp (refine, [thunk (of_variable id)]) + CAst.make ?loc @@ CTacApp (refine, [thunk (of_variable id)]) let of_dispatch tacs = - let (loc, _) = tacs in + let loc = tacs.loc in let default = function | Some e -> thunk e - | None -> thunk (Loc.tag ?loc @@ CTacCst (AbsKn (Tuple 0))) + | None -> thunk (CAst.make ?loc @@ CTacCst (AbsKn (Tuple 0))) in - let map e = of_pair default (fun l -> of_list ?loc default l) (Loc.tag ?loc e) in + let map e = of_pair default (fun l -> of_list ?loc default l) (CAst.make ?loc e) in of_pair (fun l -> of_list ?loc default l) (fun r -> of_option ?loc map r) tacs let make_red_flag l = let open Genredexpr in let rec add_flag red = function | [] -> red - | (_, flag) :: lf -> + | {v=flag} :: lf -> let red = match flag with | QBeta -> { red with rBeta = true } | QMatch -> { red with rMatch = true } | QFix -> { red with rFix = true } | QCofix -> { red with rCofix = true } | QZeta -> { red with rZeta = true } - | QConst (loc, l) -> + | QConst {loc;v=l} -> if red.rDelta then CErrors.user_err ?loc Pp.(str "Cannot set both constants to unfold and constants not to unfold"); { red with rConst = red.rConst @ l } - | QDeltaBut (loc, l) -> + | QDeltaBut {loc;v=l} -> if red.rConst <> [] && not red.rDelta then CErrors.user_err ?loc Pp.(str "Cannot set both constants to unfold and constants not to unfold"); @@ -348,10 +349,10 @@ let of_reference r = in of_anti of_ref r -let of_strategy_flag (loc, flag) = +let of_strategy_flag {loc;v=flag} = let open Genredexpr in let flag = make_red_flag flag in - Loc.tag ?loc @@ CTacRec ([ + CAst.make ?loc @@ CTacRec ([ std_proj "rBeta", of_bool ?loc flag.rBeta; std_proj "rMatch", of_bool ?loc flag.rMatch; std_proj "rFix", of_bool ?loc flag.rFix; @@ -361,7 +362,7 @@ let of_strategy_flag (loc, flag) = std_proj "rConst", of_list ?loc of_reference flag.rConst; ]) -let of_hintdb (loc, hdb) = match hdb with +let of_hintdb {loc;v=hdb} = match hdb with | QHintAll -> of_option ?loc (fun l -> of_list (fun id -> of_anti of_ident id) l) None | QHintDbs ids -> of_option ?loc (fun l -> of_list (fun id -> of_anti of_ident id) l) (Some ids) @@ -375,8 +376,8 @@ let extract_name ?loc oid = match oid with [(match_kind * pattern * (context -> constr array -> 'a))] where the function binds the names from the pattern to the contents of the constr array. *) -let of_constr_matching (loc, m) = - let map (loc, ((ploc, pat), tac)) = +let of_constr_matching {loc;v=m} = + let map {loc;v=({loc=ploc;v=pat}, tac)} = let (knd, pat, na) = match pat with | QConstrMatchPattern pat -> let knd = constructor ?loc (pattern_core "MatchPattern") [] in @@ -391,7 +392,7 @@ let of_constr_matching (loc, m) = let vars = Id.Set.elements vars in let vars = List.map (fun id -> Name id) vars in let e = abstract_vars loc vars tac in - let e = Loc.tag ?loc @@ CTacFun ([Loc.tag ?loc @@ CPatVar na], e) in + let e = CAst.make ?loc @@ CTacFun ([CAst.make ?loc @@ CPatVar na], e) in let pat = inj_wit ?loc:ploc wit_pattern pat in of_tuple [knd; pat; e] in @@ -401,8 +402,8 @@ let of_constr_matching (loc, m) = - a goal pattern: (constr_match list * constr_match) - a branch function (ident array -> context array -> constr array -> context -> 'a) *) -let of_goal_matching (loc, gm) = - let mk_pat (loc, p) = match p with +let of_goal_matching {loc;v=gm} = + let mk_pat {loc;v=p} = match p with | QConstrMatchPattern pat -> let knd = constructor ?loc (pattern_core "MatchPattern") [] in (Anonymous, pat, knd) @@ -411,7 +412,7 @@ let of_goal_matching (loc, gm) = let knd = constructor ?loc (pattern_core "MatchContext") [] in (na, pat, knd) in - let mk_gpat (loc, p) = + let mk_gpat {loc;v=p} = let concl_pat = p.q_goal_match_concl in let hyps_pats = p.q_goal_match_hyps in let (concl_ctx, concl_pat, concl_knd) = mk_pat concl_pat in @@ -433,9 +434,9 @@ let of_goal_matching (loc, gm) = let subst = List.map (fun id -> Name id) vars in (r, hyps, hctx, subst, concl_ctx) in - let map (loc, (pat, tac)) = + let map {loc;v=(pat, tac)} = let (pat, hyps, hctx, subst, cctx) = mk_gpat pat in - let tac = Loc.tag ?loc @@ CTacFun ([Loc.tag ?loc @@ CPatVar cctx], tac) in + let tac = CAst.make ?loc @@ CTacFun ([CAst.make ?loc @@ CPatVar cctx], tac) in let tac = abstract_vars loc subst tac in let tac = abstract_vars loc hctx tac in let tac = abstract_vars loc hyps tac in @@ -443,7 +444,7 @@ let of_goal_matching (loc, gm) = in of_list ?loc map gm -let of_move_location (loc, mv) = match mv with +let of_move_location {loc;v=mv} = match mv with | QMoveAfter id -> std_constructor ?loc "MoveAfter" [of_anti of_ident id] | QMoveBefore id -> std_constructor ?loc "MoveBefore" [of_anti of_ident id] | QMoveFirst -> std_constructor ?loc "MoveFirst" [] @@ -452,7 +453,7 @@ let of_move_location (loc, mv) = match mv with let of_pose p = of_pair (fun id -> of_option (fun id -> of_anti of_ident id) id) of_open_constr p -let of_assertion (loc, ast) = match ast with +let of_assertion {loc;v=ast} = match ast with | QAssertType (ipat, c, tac) -> let ipat = of_option of_intro_pattern ipat in let c = of_constr c in diff --git a/src/tac2quote.mli b/src/tac2quote.mli index 3f6c9a55e5..94e1734222 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Loc open Names open Tac2dyn open Tac2qexpr @@ -23,15 +22,15 @@ val thunk : raw_tacexpr -> raw_tacexpr val of_anti : ('a -> raw_tacexpr) -> 'a or_anti -> raw_tacexpr -val of_int : int located -> raw_tacexpr +val of_int : int CAst.t -> raw_tacexpr -val of_pair : ('a -> raw_tacexpr) -> ('b -> raw_tacexpr) -> ('a * 'b) located -> raw_tacexpr +val of_pair : ('a -> raw_tacexpr) -> ('b -> raw_tacexpr) -> ('a * 'b) CAst.t -> raw_tacexpr val of_tuple : ?loc:Loc.t -> raw_tacexpr list -> raw_tacexpr -val of_variable : Id.t located -> raw_tacexpr +val of_variable : Id.t CAst.t -> raw_tacexpr -val of_ident : Id.t located -> raw_tacexpr +val of_ident : Id.t CAst.t -> raw_tacexpr val of_constr : Constrexpr.constr_expr -> raw_tacexpr @@ -43,7 +42,7 @@ val of_bindings : bindings -> raw_tacexpr val of_intro_pattern : intro_pattern -> raw_tacexpr -val of_intro_patterns : intro_pattern list located -> raw_tacexpr +val of_intro_patterns : intro_pattern list CAst.t -> raw_tacexpr val of_clause : clause -> raw_tacexpr @@ -63,13 +62,13 @@ val of_move_location : move_location -> raw_tacexpr val of_reference : Libnames.reference or_anti -> raw_tacexpr -val of_hyp : ?loc:Loc.t -> Id.t located -> raw_tacexpr +val of_hyp : ?loc:Loc.t -> Id.t CAst.t -> raw_tacexpr (** id ↦ 'Control.hyp @id' *) -val of_exact_hyp : ?loc:Loc.t -> Id.t located -> raw_tacexpr +val of_exact_hyp : ?loc:Loc.t -> Id.t CAst.t -> raw_tacexpr (** id ↦ 'Control.refine (fun () => Control.hyp @id') *) -val of_exact_var : ?loc:Loc.t -> Id.t located -> raw_tacexpr +val of_exact_var : ?loc:Loc.t -> Id.t CAst.t -> raw_tacexpr (** id ↦ 'Control.refine (fun () => Control.hyp @id') *) val of_dispatch : dispatch -> raw_tacexpr diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index 65cdef0f3f..9bc23e91d7 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -37,16 +37,16 @@ let delayed_of_thunk r tac env sigma = let mk_bindings = function | ImplicitBindings l -> Misctypes.ImplicitBindings l | ExplicitBindings l -> - let l = List.map Loc.tag l in + let l = List.map CAst.make l in Misctypes.ExplicitBindings l | NoBindings -> Misctypes.NoBindings let mk_with_bindings (x, b) = (x, mk_bindings b) let rec mk_intro_pattern = function -| IntroForthcoming b -> Loc.tag @@ Misctypes.IntroForthcoming b -| IntroNaming ipat -> Loc.tag @@ Misctypes.IntroNaming (mk_intro_pattern_naming ipat) -| IntroAction ipat -> Loc.tag @@ Misctypes.IntroAction (mk_intro_pattern_action ipat) +| IntroForthcoming b -> CAst.make @@ Misctypes.IntroForthcoming b +| IntroNaming ipat -> CAst.make @@ Misctypes.IntroNaming (mk_intro_pattern_naming ipat) +| IntroAction ipat -> CAst.make @@ Misctypes.IntroAction (mk_intro_pattern_action ipat) and mk_intro_pattern_naming = function | IntroIdentifier id -> Misctypes.IntroIdentifier id @@ -58,7 +58,7 @@ and mk_intro_pattern_action = function | IntroOrAndPattern ipat -> Misctypes.IntroOrAndPattern (mk_or_and_intro_pattern ipat) | IntroInjection ipats -> Misctypes.IntroInjection (List.map mk_intro_pattern ipats) | IntroApplyOn (c, ipat) -> - let c = Loc.tag @@ delayed_of_thunk Tac2ffi.constr c in + let c = CAst.make @@ delayed_of_thunk Tac2ffi.constr c in Misctypes.IntroApplyOn (c, mk_intro_pattern ipat) | IntroRewrite b -> Misctypes.IntroRewrite b @@ -94,7 +94,7 @@ let intros_patterns ev ipat = let apply adv ev cb cl = let map c = let c = thaw constr_with_bindings c >>= fun p -> return (mk_with_bindings p) in - None, Loc.tag (delayed_of_tactic c) + None, CAst.make (delayed_of_tactic c) in let cb = List.map map cb in match cl with @@ -111,8 +111,8 @@ let mk_destruction_arg = function | ElimOnAnonHyp n -> Misctypes.ElimOnAnonHyp n let mk_induction_clause (arg, eqn, as_, occ) = - let eqn = Option.map (fun ipat -> Loc.tag @@ mk_intro_pattern_naming ipat) eqn in - let as_ = Option.map (fun ipat -> Loc.tag @@ mk_or_and_intro_pattern ipat) as_ in + let eqn = Option.map (fun ipat -> CAst.make @@ mk_intro_pattern_naming ipat) eqn in + let as_ = Option.map (fun ipat -> CAst.make @@ mk_or_and_intro_pattern ipat) as_ in let occ = Option.map mk_clause occ in ((None, mk_destruction_arg arg), (eqn, as_), occ) @@ -188,7 +188,7 @@ let forward fst tac ipat c = let assert_ = function | AssertValue (id, c) -> - let ipat = Loc.tag @@ Misctypes.IntroNaming (Misctypes.IntroIdentifier id) in + let ipat = CAst.make @@ Misctypes.IntroNaming (Misctypes.IntroIdentifier id) in Tactics.forward true None (Some ipat) c | AssertType (ipat, c, tac) -> let ipat = Option.map mk_intro_pattern ipat in @@ -196,7 +196,7 @@ let assert_ = function Tactics.forward true (Some tac) ipat c let letin_pat_tac ev ipat na c cl = - let ipat = Option.map (fun (b, ipat) -> (b, Loc.tag @@ mk_intro_pattern_naming ipat)) ipat in + let ipat = Option.map (fun (b, ipat) -> (b, CAst.make @@ mk_intro_pattern_naming ipat)) ipat in let cl = mk_clause cl in Tactics.letin_pat_tac ev ipat na c cl @@ -420,7 +420,7 @@ let inversion knd arg pat ids = begin match pat with | None -> Proofview.tclUNIT None | Some (IntroAction (IntroOrAndPattern p)) -> - Proofview.tclUNIT (Some (Loc.tag @@ mk_or_and_intro_pattern p)) + Proofview.tclUNIT (Some (CAst.make @@ mk_or_and_intro_pattern p)) | Some _ -> Tacticals.New.tclZEROMSG (str "Inversion only accept disjunctive patterns") end >>= fun pat -> @@ -433,7 +433,7 @@ let inversion knd arg pat ids = Inv.inv_clause knd pat ids (NamedHyp id) | Some (_, Misctypes.ElimOnConstr c) -> let open Misctypes in - let anon = Loc.tag @@ IntroNaming IntroAnonymous in + let anon = CAst.make @@ IntroNaming IntroAnonymous in Tactics.specialize c (Some anon) >>= fun () -> Tacticals.New.onLastHypId (fun id -> Inv.inv_clause knd pat ids (NamedHyp id)) end -- cgit v1.2.3 From 4718f5afec0538781195fdac82ffefb0b7c57535 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Mon, 5 Mar 2018 23:38:52 +0100 Subject: [coq] Adapt to coq/coq#6837. Minor. --- src/g_ltac2.ml4 | 8 ++++---- src/tac2core.ml | 6 +++--- src/tac2entries.ml | 2 +- src/tac2quote.ml | 3 +-- 4 files changed, 9 insertions(+), 10 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index f4818f4ece..6189bb18cc 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -372,8 +372,8 @@ GEXTEND Gram [ [ id = Prim.ident -> CAst.make ~loc:!@loc id ] ] ; globref: - [ [ "&"; id = Prim.ident -> Libnames.Ident (Loc.tag ~loc:!@loc id) - | qid = Prim.qualid -> Libnames.Qualid (Loc.tag ~loc:!@loc qid.CAst.v) + [ [ "&"; id = Prim.ident -> CAst.make ~loc:!@loc (Libnames.Ident id) + | qid = Prim.qualid -> CAst.map (fun qid -> Libnames.Qualid qid) qid ] ] ; END @@ -667,8 +667,8 @@ GEXTEND Gram ] ] ; refglobal: - [ [ "&"; id = Prim.ident -> QExpr (Libnames.Ident (Loc.tag ~loc:!@loc id)) - | qid = Prim.qualid -> QExpr (Libnames.Qualid Loc.(tag ~loc:!@loc qid.CAst.v)) + [ [ "&"; id = Prim.ident -> QExpr (CAst.make ~loc:!@loc @@ Libnames.Ident id) + | qid = Prim.qualid -> QExpr (CAst.make ~loc:!@loc @@ Libnames.Qualid qid.CAst.v) | "$"; id = Prim.ident -> QAnti (CAst.make ~loc:!@loc id) ] ] ; diff --git a/src/tac2core.ml b/src/tac2core.ml index c6c3f26b6b..cf9af526d4 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -910,13 +910,13 @@ let () = let () = let intern self ist qid = match qid with - | Libnames.Ident (_, id) -> + | {CAst.v=Libnames.Ident id} -> GlbVal (Globnames.VarRef id), gtypref t_reference - | Libnames.Qualid (loc, qid) -> + | {CAst.loc;v=Libnames.Qualid qid} -> let gr = try Nametab.locate qid with Not_found -> - Nametab.error_global_not_found ?loc qid + Nametab.error_global_not_found (CAst.make ?loc qid) in GlbVal gr, gtypref t_reference in diff --git a/src/tac2entries.ml b/src/tac2entries.ml index e4bddf439b..430142832b 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -831,7 +831,7 @@ end (** Printing *) let print_ltac ref = - let (loc, qid) = qualid_of_reference ref in + let {loc;v=qid} = qualid_of_reference ref in if Tac2env.is_constructor qid then let kn = try Tac2env.locate_constructor qid diff --git a/src/tac2quote.ml b/src/tac2quote.ml index e986bfc393..2a0230b779 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -344,8 +344,7 @@ let make_red_flag l = let of_reference r = let of_ref ref = - let loc = Libnames.loc_of_reference ref in - inj_wit ?loc wit_reference ref + inj_wit ?loc:ref.loc wit_reference ref in of_anti of_ref r -- cgit v1.2.3 From 7dbf6dbd73f3a828bbb58458191912f895cf7779 Mon Sep 17 00:00:00 2001 From: Armael Date: Tue, 13 Mar 2018 16:13:18 +0100 Subject: Fix a typo --- doc/ltac2.md | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/doc/ltac2.md b/doc/ltac2.md index cd0d8f4325..10a65fce44 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -875,7 +875,7 @@ Due to conflicts, a few syntactic rules have changed. `try`, `repeat`, `do`, `once`, `progress`, `time`, `abstract`. - `idtac` is no more. Either use `()` if you expect nothing to happen, `(fun () => ())` if you want a thunk (see next section), or use printing - primitives from the `Message` module if you wand to display something. + primitives from the `Message` module if you want to display something. ## Tactic delay -- cgit v1.2.3 From e31fa0ecb1f823cc5735fa63330f430e2c0d96e2 Mon Sep 17 00:00:00 2001 From: Armael Date: Tue, 13 Mar 2018 17:16:55 +0100 Subject: Fix another typo in the documentation's grammar for open variants --- doc/ltac2.md | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/doc/ltac2.md b/doc/ltac2.md index 10a65fce44..9ba227c285 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -127,7 +127,7 @@ TYPEDEF := | TYPE | "[" CONSTRUCTORDEF₀ "|" ... "|" CONSTRUCTORDEFₙ "]" | "{" FIELDDEF₀ ";" ... ";" FIELDDEFₙ "}" -| "[" "..." "]" +| "[" ".." "]" CONSTRUCTORDEF := | IDENT ( "(" TYPE₀ "," ... "," TYPE₀ ")" ) -- cgit v1.2.3 From 33d7e374b449c754fcdf623e98cb717faaf241a5 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 4 Dec 2017 17:45:50 +0100 Subject: Fixup strict mode for patterns --- src/tac2core.ml | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index fe4912a6c5..32c873088c 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -895,7 +895,8 @@ let () = let intern self ist c = let env = ist.Genintern.genv in let sigma = Evd.from_env env in - let _, pat = Constrintern.intern_constr_pattern env sigma ~as_type:false c in + let warn = if !Ltac_plugin.Tacintern.strict_check then fun x -> x else Constrintern.for_grammar in + let _, pat = warn (fun () ->Constrintern.intern_constr_pattern env sigma ~as_type:false c) () in GlbVal pat, gtypref t_pattern in let print env pat = str "pattern:(" ++ Printer.pr_lconstr_pattern_env env Evd.empty pat ++ str ")" in -- cgit v1.2.3 From 92880aa90abe810115227e1e2dd67355d7f5c872 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 5 Dec 2017 13:13:11 +0100 Subject: Fix printing of toplevel record values. --- src/tac2print.ml | 15 ++++++++++++++- 1 file changed, 14 insertions(+), 1 deletion(-) diff --git a/src/tac2print.ml b/src/tac2print.ml index 8f61686988..cab1530d15 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -390,7 +390,9 @@ let rec pr_valexpr env sigma v t = match kind t with let knc = change_kn_label kn id in let args = pr_constrargs env sigma params args tpe in hv 2 (pr_constructor knc ++ spc () ++ str "(" ++ args ++ str ")") - | GTydRec rcd -> str "{ TODO }" + | GTydRec rcd -> + let (_, args) = Tac2ffi.to_block v in + pr_record env sigma params args rcd | GTydOpn -> begin match Tac2ffi.to_open v with | (knc, [||]) -> pr_constructor knc @@ -417,6 +419,17 @@ and pr_constrargs env sigma params args tpe = let args = List.combine args tpe in prlist_with_sep pr_comma (fun (v, t) -> pr_valexpr env sigma v t) args +and pr_record env sigma params args rcd = + let subst = Array.of_list params in + let map (id, _, tpe) = (id, subst_type subst tpe) in + let rcd = List.map map rcd in + let args = Array.to_list args in + let fields = List.combine rcd args in + let pr_field ((id, t), arg) = + Id.print id ++ spc () ++ str ":=" ++ spc () ++ pr_valexpr env sigma arg t + in + str "{" ++ spc () ++ prlist_with_sep pr_semicolon pr_field fields ++ spc () ++ str "}" + and pr_val_list env sigma args tpe = let pr v = pr_valexpr env sigma v tpe in str "[" ++ prlist_with_sep pr_semicolon pr args ++ str "]" -- cgit v1.2.3 From 15a9b4a7a99498895addb74ffa9a711ea354c651 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Mon, 2 Apr 2018 03:51:52 +0200 Subject: [coq] Adapt to coq/coq#6960. Minor, a couple of tactic-related types moved. --- src/tac2stdlib.ml | 6 +++--- src/tac2tactics.ml | 18 +++++++++--------- src/tac2tactics.mli | 2 +- 3 files changed, 13 insertions(+), 13 deletions(-) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 28d4967874..9ed3d5b32e 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -186,9 +186,9 @@ let to_strategy v = match Value.to_int v with let strategy = make_to_repr to_strategy let to_inversion_kind v = match Value.to_int v with -| 0 -> Misctypes.SimpleInversion -| 1 -> Misctypes.FullInversion -| 2 -> Misctypes.FullInversionClear +| 0 -> Inv.SimpleInversion +| 1 -> Inv.FullInversion +| 2 -> Inv.FullInversionClear | _ -> assert false let inversion_kind = make_to_repr to_inversion_kind diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index 9bc23e91d7..3f2385a8d6 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -106,9 +106,9 @@ let apply adv ev cb cl = let mk_destruction_arg = function | ElimOnConstr c -> let c = c >>= fun c -> return (mk_with_bindings c) in - Misctypes.ElimOnConstr (delayed_of_tactic c) -| ElimOnIdent id -> Misctypes.ElimOnIdent CAst.(make id) -| ElimOnAnonHyp n -> Misctypes.ElimOnAnonHyp n + Tactics.ElimOnConstr (delayed_of_tactic c) +| ElimOnIdent id -> Tactics.ElimOnIdent CAst.(make id) +| ElimOnAnonHyp n -> Tactics.ElimOnAnonHyp n let mk_induction_clause (arg, eqn, as_, occ) = let eqn = Option.map (fun ipat -> CAst.make @@ mk_intro_pattern_naming ipat) eqn in @@ -342,9 +342,9 @@ let on_destruction_arg tac ev arg = Proofview.tclEVARMAP >>= fun sigma' -> let flags = tactic_infer_flags ev in let (sigma', c) = Unification.finish_evar_resolution ~flags env sigma' (sigma, c) in - Proofview.tclUNIT (Some sigma', Misctypes.ElimOnConstr (c, lbind)) - | ElimOnIdent id -> Proofview.tclUNIT (None, Misctypes.ElimOnIdent CAst.(make id)) - | ElimOnAnonHyp n -> Proofview.tclUNIT (None, Misctypes.ElimOnAnonHyp n) + Proofview.tclUNIT (Some sigma', Tactics.ElimOnConstr (c, lbind)) + | ElimOnIdent id -> Proofview.tclUNIT (None, Tactics.ElimOnIdent CAst.(make id)) + | ElimOnAnonHyp n -> Proofview.tclUNIT (None, Tactics.ElimOnAnonHyp n) in arg >>= fun (sigma', arg) -> let arg = Some (clear, arg) in @@ -427,11 +427,11 @@ let inversion knd arg pat ids = let inversion _ arg = begin match arg with | None -> assert false - | Some (_, Misctypes.ElimOnAnonHyp n) -> + | Some (_, Tactics.ElimOnAnonHyp n) -> Inv.inv_clause knd pat ids (AnonHyp n) - | Some (_, Misctypes.ElimOnIdent {CAst.v=id}) -> + | Some (_, Tactics.ElimOnIdent {CAst.v=id}) -> Inv.inv_clause knd pat ids (NamedHyp id) - | Some (_, Misctypes.ElimOnConstr c) -> + | Some (_, Tactics.ElimOnConstr c) -> let open Misctypes in let anon = CAst.make @@ IntroNaming IntroAnonymous in Tactics.specialize c (Some anon) >>= fun () -> diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli index 52e8a94c19..631e36b5ae 100644 --- a/src/tac2tactics.mli +++ b/src/tac2tactics.mli @@ -118,7 +118,7 @@ val eauto : Hints.debug -> int option -> int option -> constr thunk list -> val typeclasses_eauto : Class_tactics.search_strategy option -> int option -> Id.t list option -> unit Proofview.tactic -val inversion : Misctypes.inversion_kind -> destruction_arg -> intro_pattern option -> Id.t list option -> unit tactic +val inversion : Inv.inversion_kind -> destruction_arg -> intro_pattern option -> Id.t list option -> unit tactic val contradiction : constr_with_bindings option -> unit tactic -- cgit v1.2.3 From a255ba63f82652ac7dae4a39e80c35f813d0e5a3 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 11 Apr 2018 08:30:56 +0200 Subject: Fix compilation w.r.t. coq/coq#7213. --- src/tac2core.ml | 4 ++-- src/tac2match.ml | 2 +- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 32c873088c..9a3aed3442 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -557,7 +557,7 @@ let () = define2 "pattern_matches_subterm" pattern constr begin fun pat c -> | IStream.Cons ({ m_sub = (_, sub); m_ctx }, s) -> let ans = Id.Map.bindings sub in let of_pair (id, c) = Value.of_tuple [| Value.of_ident id; Value.of_constr c |] in - let ans = Value.of_tuple [| Value.of_constr m_ctx; Value.of_list of_pair ans |] in + let ans = Value.of_tuple [| Value.of_constr (Lazy.force m_ctx); Value.of_list of_pair ans |] in Proofview.tclOR (return ans) (fun _ -> of_ans s) in pf_apply begin fun env sigma -> @@ -589,7 +589,7 @@ let () = define2 "pattern_matches_subterm_vect" pattern constr begin fun pat c - | IStream.Cons ({ m_sub = (_, sub); m_ctx }, s) -> let ans = Id.Map.bindings sub in let ans = Array.map_of_list snd ans in - let ans = Value.of_tuple [| Value.of_constr m_ctx; Value.of_array Value.of_constr ans |] in + let ans = Value.of_tuple [| Value.of_constr (Lazy.force m_ctx); Value.of_array Value.of_constr ans |] in Proofview.tclOR (return ans) (fun _ -> of_ans s) in pf_apply begin fun env sigma -> diff --git a/src/tac2match.ml b/src/tac2match.ml index 5035c9dba6..a3140eabea 100644 --- a/src/tac2match.ml +++ b/src/tac2match.ml @@ -170,7 +170,7 @@ module PatternMatching (E:StaticEnvironment) = struct let nctx = { subst } in match merge ctx nctx with | None -> (map s (e, info)).stream k ctx - | Some nctx -> Proofview.tclOR (k (Some m_ctx) nctx) (fun e -> (map s e).stream k ctx) + | Some nctx -> Proofview.tclOR (k (Some (Lazy.force m_ctx)) nctx) (fun e -> (map s e).stream k ctx) } in map (Constr_matching.match_subterm E.env E.sigma (Id.Set.empty,p) term) imatching_error -- cgit v1.2.3 From 513884e806a4db39ae6402333833ecc4f70a0fdc Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Thu, 19 Apr 2018 08:00:43 +0200 Subject: Fixing printing level for subtypes of a type constructor. --- src/tac2print.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/tac2print.ml b/src/tac2print.ml index cab1530d15..9c530dfc51 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -49,7 +49,7 @@ let pr_glbtype_gen pr lvl c = | T5_r | T5_l | T2 | T1 -> fun x -> x | T0 -> paren in - paren (pr_glbtype lvl t ++ spc () ++ pr_typref kn) + paren (pr_glbtype T1 t ++ spc () ++ pr_typref kn) | GTypRef (Other kn, tl) -> let paren = match lvl with | T5_r | T5_l | T2 | T1 -> fun x -> x -- cgit v1.2.3 From 5f9df6d084bb24ea9b26a74387b79656e4123ee0 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Thu, 19 Apr 2018 08:05:31 +0200 Subject: Allowing formatting break around a printed type. --- src/tac2intern.ml | 16 ++++++++-------- 1 file changed, 8 insertions(+), 8 deletions(-) diff --git a/src/tac2intern.ml b/src/tac2intern.ml index b1dd8f7f51..99965c4ed5 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -356,8 +356,8 @@ let rec unify0 env t1 t2 = match kind env t1, kind env t2 with let unify ?loc env t1 t2 = try unify0 env t1 t2 with CannotUnify (u1, u2) -> - user_err ?loc (str "This expression has type " ++ pr_glbtype env t1 ++ - str " but an expression was expected of type " ++ pr_glbtype env t2) + user_err ?loc (str "This expression has type" ++ spc () ++ pr_glbtype env t1 ++ + spc () ++ str "but an expression was expected of type" ++ spc () ++ pr_glbtype env t2) let unify_arrow ?loc env ft args = let ft0 = ft in @@ -372,11 +372,11 @@ let unify_arrow ?loc env ft args = iter ft args true | GTypRef _, _ :: _ -> if is_fun then - user_err ?loc (str "This function has type " ++ pr_glbtype env ft0 ++ - str " and is applied to too many arguments") + user_err ?loc (str "This function has type" ++ spc () ++ pr_glbtype env ft0 ++ + spc () ++ str "and is applied to too many arguments") else - user_err ?loc (str "This expression has type " ++ pr_glbtype env ft0 ++ - str " and is not a function") + user_err ?loc (str "This expression has type" ++ spc () ++ pr_glbtype env ft0 ++ + spc () ++ str "and is not a function") in iter ft args false @@ -475,13 +475,13 @@ let check_elt_empty loc env t = match kind env t with | GTypVar _ -> user_err ?loc (str "Cannot infer an empty type for this expression") | GTypArrow _ | GTypRef (Tuple _, _) -> - user_err ?loc (str "Type " ++ pr_glbtype env t ++ str " is not an empty type") + user_err ?loc (str "Type" ++ spc () ++ pr_glbtype env t ++ spc () ++ str "is not an empty type") | GTypRef (Other kn, _) -> let def = Tac2env.interp_type kn in match def with | _, GTydAlg { galg_constructors = [] } -> kn | _ -> - user_err ?loc (str "Type " ++ pr_glbtype env t ++ str " is not an empty type") + user_err ?loc (str "Type" ++ spc () ++ pr_glbtype env t ++ spc () ++ str "is not an empty type") let check_unit ?loc t = let env = empty_env () in -- cgit v1.2.3 From 8db84eef2cf6a5f35e9ef7e4bea9725f29550845 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Mon, 14 May 2018 03:59:55 +0200 Subject: Adapt to fix/cofix changes in Coq (coq/coq#7196) c.f. https://github.com/coq/coq/pull/7196 --- src/tac2stdlib.ml | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index 9ed3d5b32e..be7c76744d 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -501,12 +501,12 @@ end let () = define_prim0 "tac_admit" Proofview.give_up -let () = define_prim2 "tac_fix" (option ident) int begin fun idopt n -> - Tactics.fix idopt n +let () = define_prim2 "tac_fix" ident int begin fun ident n -> + Tactics.fix ident n end -let () = define_prim1 "tac_cofix" (option ident) begin fun idopt -> - Tactics.cofix idopt +let () = define_prim1 "tac_cofix" ident begin fun ident -> + Tactics.cofix ident end let () = define_prim1 "tac_clear" (list ident) begin fun ids -> -- cgit v1.2.3 From 5eb92624f33afb4d2a390f7c7b80552e3e04bc81 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Fri, 25 May 2018 16:03:01 +0200 Subject: Renaming smartmap and fold_map according to Coq PR #7177. --- src/tac2entries.ml | 4 ++-- src/tac2intern.ml | 30 +++++++++++++++--------------- 2 files changed, 17 insertions(+), 17 deletions(-) diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 430142832b..8728a513cf 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -248,12 +248,12 @@ let open_typext i obj = perform_typext (Exactly i) obj let subst_typext (subst, e) = let open Mod_subst in let subst_data data = - let edata_args = List.smartmap (fun e -> subst_type subst e) data.edata_args in + let edata_args = List.Smart.map (fun e -> subst_type subst e) data.edata_args in if edata_args == data.edata_args then data else { data with edata_args } in let typext_type = subst_kn subst e.typext_type in - let typext_expr = List.smartmap subst_data e.typext_expr in + let typext_expr = List.Smart.map subst_data e.typext_expr in if typext_type == e.typext_type && typext_expr == e.typext_expr then e else diff --git a/src/tac2intern.ml b/src/tac2intern.ml index b1dd8f7f51..75fca938f4 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -840,7 +840,7 @@ and intern_let_rec env loc ids el e = let env = push_name na (monomorphic (GTypVar id)) env in (env, (loc, na, t, e, id)) in - let (env, el) = List.fold_map map env el in + let (env, el) = List.fold_left_map map env el in let fold (loc, na, tc, e, id) (el, tl) = let loc_e = e.loc in let (e, t) = intern_rec env e in @@ -1314,7 +1314,7 @@ let rec subst_type subst t = match t with else GTypArrow (t1', t2') | GTypRef (kn, tl) -> let kn' = subst_or_tuple subst_kn subst kn in - let tl' = List.smartmap (fun t -> subst_type subst t) tl in + let tl' = List.Smart.map (fun t -> subst_type subst t) tl in if kn' == kn && tl' == tl then t else GTypRef (kn', tl') let rec subst_expr subst e = match e with @@ -1328,7 +1328,7 @@ let rec subst_expr subst e = match e with GTacLet (r, bs, subst_expr subst e) | GTacCst (t, n, el) as e0 -> let t' = subst_or_tuple subst_kn subst t in - let el' = List.smartmap (fun e -> subst_expr subst e) el in + let el' = List.Smart.map (fun e -> subst_expr subst e) el in if t' == t && el' == el then e0 else GTacCst (t', n, el') | GTacCse (e, ci, cse0, cse1) -> let cse0' = Array.map (fun e -> subst_expr subst e) cse0 in @@ -1362,19 +1362,19 @@ let rec subst_expr subst e = match e with if arg' == arg then e else GTacExt (tag, arg') | GTacOpn (kn, el) as e0 -> let kn' = subst_kn subst kn in - let el' = List.smartmap (fun e -> subst_expr subst e) el in + let el' = List.Smart.map (fun e -> subst_expr subst e) el in if kn' == kn && el' == el then e0 else GTacOpn (kn', el') let subst_typedef subst e = match e with | GTydDef t -> - let t' = Option.smartmap (fun t -> subst_type subst t) t in + let t' = Option.Smart.map (fun t -> subst_type subst t) t in if t' == t then e else GTydDef t' | GTydAlg galg -> let map (c, tl as p) = - let tl' = List.smartmap (fun t -> subst_type subst t) tl in + let tl' = List.Smart.map (fun t -> subst_type subst t) tl in if tl' == tl then p else (c, tl') in - let constrs' = List.smartmap map galg.galg_constructors in + let constrs' = List.Smart.map map galg.galg_constructors in if constrs' == galg.galg_constructors then e else GTydAlg { galg with galg_constructors = constrs' } | GTydRec fields -> @@ -1382,7 +1382,7 @@ let subst_typedef subst e = match e with let t' = subst_type subst t in if t' == t then p else (c, mut, t') in - let fields' = List.smartmap map fields in + let fields' = List.Smart.map map fields in if fields' == fields then e else GTydRec fields' | GTydOpn -> GTydOpn @@ -1408,7 +1408,7 @@ let rec subst_rawtype subst ({loc;v=tr} as t) = match tr with if t1' == t1 && t2' == t2 then t else CAst.make ?loc @@ CTypArrow (t1', t2') | CTypRef (ref, tl) -> let ref' = subst_or_relid subst ref in - let tl' = List.smartmap (fun t -> subst_rawtype subst t) tl in + let tl' = List.Smart.map (fun t -> subst_rawtype subst t) tl in if ref' == ref && tl' == tl then t else CAst.make ?loc @@ CTypRef (ref', tl') let subst_tacref subst ref = match ref with @@ -1429,7 +1429,7 @@ let subst_projection subst prj = match prj with let rec subst_rawpattern subst ({loc;v=pr} as p) = match pr with | CPatVar _ -> p | CPatRef (c, pl) -> - let pl' = List.smartmap (fun p -> subst_rawpattern subst p) pl in + let pl' = List.Smart.map (fun p -> subst_rawpattern subst p) pl in let c' = subst_or_relid subst c in if pl' == pl && c' == c then p else CAst.make ?loc @@ CPatRef (c', pl') | CPatCnv (pat, ty) -> @@ -1448,12 +1448,12 @@ let rec subst_rawexpr subst ({loc;v=tr} as t) = match tr with if ref' == ref then t else CAst.make ?loc @@ CTacCst ref' | CTacFun (bnd, e) -> let map pat = subst_rawpattern subst pat in - let bnd' = List.smartmap map bnd in + let bnd' = List.Smart.map map bnd in let e' = subst_rawexpr subst e in if bnd' == bnd && e' == e then t else CAst.make ?loc @@ CTacFun (bnd', e') | CTacApp (e, el) -> let e' = subst_rawexpr subst e in - let el' = List.smartmap (fun e -> subst_rawexpr subst e) el in + let el' = List.Smart.map (fun e -> subst_rawexpr subst e) el in if e' == e && el' == el then t else CAst.make ?loc @@ CTacApp (e', el') | CTacLet (isrec, bnd, e) -> let map (na, e as p) = @@ -1461,7 +1461,7 @@ let rec subst_rawexpr subst ({loc;v=tr} as t) = match tr with let e' = subst_rawexpr subst e in if na' == na && e' == e then p else (na', e') in - let bnd' = List.smartmap map bnd in + let bnd' = List.Smart.map map bnd in let e' = subst_rawexpr subst e in if bnd' == bnd && e' == e then t else CAst.make ?loc @@ CTacLet (isrec, bnd', e') | CTacCnv (e, c) -> @@ -1479,7 +1479,7 @@ let rec subst_rawexpr subst ({loc;v=tr} as t) = match tr with if p' == p && e' == e then x else (p', e') in let e' = subst_rawexpr subst e in - let bl' = List.smartmap map bl in + let bl' = List.Smart.map map bl in if e' == e && bl' == bl then t else CAst.make ?loc @@ CTacCse (e', bl') | CTacRec el -> let map (prj, e as p) = @@ -1487,7 +1487,7 @@ let rec subst_rawexpr subst ({loc;v=tr} as t) = match tr with let e' = subst_rawexpr subst e in if prj' == prj && e' == e then p else (prj', e') in - let el' = List.smartmap map el in + let el' = List.Smart.map map el in if el' == el then t else CAst.make ?loc @@ CTacRec el' | CTacPrj (e, prj) -> let prj' = subst_projection subst prj in -- cgit v1.2.3 From e15c228709351f5001f6a10765a816cebcf900cc Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Fri, 25 May 2018 16:04:18 +0200 Subject: Renaming global_reference according to Coq PR #6156. --- src/tac2ffi.mli | 6 +++--- src/tac2quote.mli | 2 +- src/tac2tactics.mli | 23 +++++++++++------------ 3 files changed, 15 insertions(+), 16 deletions(-) diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index 1bf86b516a..d801c4f605 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -131,9 +131,9 @@ val of_constant : Constant.t -> valexpr val to_constant : valexpr -> Constant.t val constant : Constant.t repr -val of_reference : Globnames.global_reference -> valexpr -val to_reference : valexpr -> Globnames.global_reference -val reference : Globnames.global_reference repr +val of_reference : GlobRef.t -> valexpr +val to_reference : valexpr -> GlobRef.t +val reference : GlobRef.t repr val of_ext : 'a Val.tag -> 'a -> valexpr val to_ext : 'a Val.tag -> valexpr -> 'a diff --git a/src/tac2quote.mli b/src/tac2quote.mli index 18533a8281..cc58144901 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -89,7 +89,7 @@ val wit_pattern : (Constrexpr.constr_expr, Pattern.constr_pattern) Arg.tag val wit_ident : (Id.t, Id.t) Arg.tag -val wit_reference : (Libnames.reference, Globnames.global_reference) Arg.tag +val wit_reference : (Libnames.reference, GlobRef.t) Arg.tag (** Beware, at the raw level, [Qualid [id]] has not the same meaning as [Ident id]. The first is an unqualified global reference, the second is the dynamic reference to id. *) diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli index 631e36b5ae..026673acbf 100644 --- a/src/tac2tactics.mli +++ b/src/tac2tactics.mli @@ -7,7 +7,6 @@ (************************************************************************) open Names -open Globnames open Tac2expr open EConstr open Genredexpr @@ -57,16 +56,16 @@ val letin_pat_tac : evars_flag -> (bool * intro_pattern_naming) option -> val reduce : Redexpr.red_expr -> clause -> unit tactic -val simpl : global_reference glob_red_flag -> +val simpl : GlobRef.t glob_red_flag -> (Pattern.constr_pattern * occurrences) option -> clause -> unit tactic -val cbv : global_reference glob_red_flag -> clause -> unit tactic +val cbv : GlobRef.t glob_red_flag -> clause -> unit tactic -val cbn : global_reference glob_red_flag -> clause -> unit tactic +val cbn : GlobRef.t glob_red_flag -> clause -> unit tactic -val lazy_ : global_reference glob_red_flag -> clause -> unit tactic +val lazy_ : GlobRef.t glob_red_flag -> clause -> unit tactic -val unfold : (global_reference * occurrences) list -> clause -> unit tactic +val unfold : (GlobRef.t * occurrences) list -> clause -> unit tactic val pattern : (constr * occurrences) list -> clause -> unit tactic @@ -78,16 +77,16 @@ val eval_red : constr -> constr tactic val eval_hnf : constr -> constr tactic -val eval_simpl : global_reference glob_red_flag -> +val eval_simpl : GlobRef.t glob_red_flag -> (Pattern.constr_pattern * occurrences) option -> constr -> constr tactic -val eval_cbv : global_reference glob_red_flag -> constr -> constr tactic +val eval_cbv : GlobRef.t glob_red_flag -> constr -> constr tactic -val eval_cbn : global_reference glob_red_flag -> constr -> constr tactic +val eval_cbn : GlobRef.t glob_red_flag -> constr -> constr tactic -val eval_lazy : global_reference glob_red_flag -> constr -> constr tactic +val eval_lazy : GlobRef.t glob_red_flag -> constr -> constr tactic -val eval_unfold : (global_reference * occurrences) list -> constr -> constr tactic +val eval_unfold : (GlobRef.t * occurrences) list -> constr -> constr tactic val eval_fold : constr list -> constr -> constr tactic @@ -122,4 +121,4 @@ val inversion : Inv.inversion_kind -> destruction_arg -> intro_pattern option -> val contradiction : constr_with_bindings option -> unit tactic -val firstorder : unit thunk option -> global_reference list -> Id.t list -> unit tactic +val firstorder : unit thunk option -> GlobRef.t list -> Id.t list -> unit tactic -- cgit v1.2.3 From 1cd62aa20ba3f9fb206f2926c8e47dd463dee863 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Sun, 20 May 2018 21:15:44 +0200 Subject: Adapt to coq/coq#7558. Trivial renaming of the vernacular parsing entry point. --- src/g_ltac2.ml4 | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 6189bb18cc..ac65495fe5 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -841,8 +841,8 @@ END let _ = let mode = { Proof_global.name = "Ltac2"; - set = (fun () -> set_command_entry tac2mode); - reset = (fun () -> set_command_entry Pcoq.Vernac_.noedit_mode); + set = (fun () -> Pvernac.set_command_entry tac2mode); + reset = (fun () -> Pvernac.(set_command_entry Vernac_.noedit_mode)); } in Proof_global.register_proof_mode mode -- cgit v1.2.3 From da142cf79f44c163db9c290947b34860ac902832 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 28 May 2018 19:16:24 +0200 Subject: Fix w.r.t. coq/coq#7521. --- src/tac2core.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 9a3aed3442..8d0ef675fe 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -502,7 +502,7 @@ let () = define3 "constr_in_context" ident constr closure begin fun id t c -> let sigma = Proofview.Goal.sigma gl in let has_var = try - let _ = Environ.lookup_named_val id (Environ.named_context_val env) in + let _ = Environ.lookup_named_val id env in true with Not_found -> false in -- cgit v1.2.3 From 1bbeba35eb385f813a0e4b6d25a437f9bab8191b Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 18 Jun 2018 14:32:58 +0200 Subject: Fixing a batch of deprecation warnings. --- src/g_ltac2.ml4 | 7 +++---- src/tac2core.ml | 9 ++++----- src/tac2entries.ml | 5 +++-- src/tac2entries.mli | 4 ++-- src/tac2expr.mli | 7 +++---- src/tac2intern.ml | 2 +- src/tac2qexpr.mli | 2 +- src/tac2stdlib.ml | 10 +++++----- src/tac2tactics.ml | 40 ++++++++++++++++++++-------------------- src/tac2types.mli | 4 ++-- 10 files changed, 44 insertions(+), 46 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index ac65495fe5..b59a5e184e 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -12,7 +12,6 @@ open Names open Tok open Pcoq open Constrexpr -open Misctypes open Tac2expr open Tac2qexpr open Ltac_plugin @@ -799,15 +798,15 @@ GEXTEND Gram Pcoq.Constr.operconstr: LEVEL "0" [ [ IDENT "ltac2"; ":"; "("; tac = tac2expr; ")" -> let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in - CAst.make ~loc:!@loc (CHole (None, IntroAnonymous, Some arg)) + CAst.make ~loc:!@loc (CHole (None, Namegen.IntroAnonymous, Some arg)) | test_ampersand_ident; "&"; id = Prim.ident -> let tac = Tac2quote.of_exact_hyp ~loc:!@loc (CAst.make ~loc:!@loc id) in let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in - CAst.make ~loc:!@loc (CHole (None, IntroAnonymous, Some arg)) + CAst.make ~loc:!@loc (CHole (None, Namegen.IntroAnonymous, Some arg)) | test_dollar_ident; "$"; id = Prim.ident -> let id = Loc.tag ~loc:!@loc id in let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2_quotation) id in - CAst.make ~loc:!@loc (CHole (None, IntroAnonymous, Some arg)) + CAst.make ~loc:!@loc (CHole (None, Namegen.IntroAnonymous, Some arg)) ] ] ; END diff --git a/src/tac2core.ml b/src/tac2core.ml index 8d0ef675fe..5f33374486 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -948,14 +948,14 @@ let () = let lfun = Tac2interp.set_env ist Id.Map.empty in let ist = Ltac_plugin.Tacinterp.default_ist () in let ist = { ist with Geninterp.lfun = lfun } in - let tac = (Obj.magic Ltac_plugin.Tacinterp.eval_tactic_ist ist tac : unit Proofview.tactic) in + let tac = (Ltac_plugin.Tacinterp.eval_tactic_ist ist tac : unit Proofview.tactic) in let wrap (e, info) = set_bt info >>= fun info -> Proofview.tclZERO ~info e in Proofview.tclOR tac wrap >>= fun () -> return v_unit in let subst s tac = Genintern.substitute Ltac_plugin.Tacarg.wit_tactic s tac in let print env tac = - str "ltac1:(" ++ Ltac_plugin.Pptactic.pr_glob_tactic (Obj.magic env) tac ++ str ")" + str "ltac1:(" ++ Ltac_plugin.Pptactic.pr_glob_tactic env tac ++ str ")" in let obj = { ml_intern = intern; @@ -981,9 +981,8 @@ let () = let ist = Tac2interp.get_env ist in let c = Id.Map.find id ist.env_ist in let c = Value.to_constr c in - let evdref = ref sigma in - let () = Typing.e_check env evdref c concl in - (c, !evdref) + let sigma = Typing.check env sigma c concl in + (c, sigma) in Pretyping.register_constr_interp0 wit_ltac2_quotation interp diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 8728a513cf..c85ffb2539 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -748,7 +748,7 @@ let perform_eval e = let v = Tac2interp.interp Tac2interp.empty_environment e in let selector, proof = try - Proof_bullet.get_default_goal_selector (), + Goal_select.get_default_goal_selector (), Proof_global.give_me_the_proof () with Proof_global.NoCurrentProof -> let sigma = Evd.from_env env in @@ -759,6 +759,7 @@ let perform_eval e = | Vernacexpr.SelectList l -> Proofview.tclFOCUSLIST l v | Vernacexpr.SelectId id -> Proofview.tclFOCUSID id v | Vernacexpr.SelectAll -> v + | Vernacexpr.SelectAlreadyFocused -> assert false (** TODO **) in (** HACK: the API doesn't allow to return a value *) let ans = ref None in @@ -864,7 +865,7 @@ let print_ltac ref = let solve default tac = let status = Proof_global.with_current_proof begin fun etac p -> let with_end_tac = if default then Some etac else None in - let g = Proof_bullet.get_default_goal_selector () in + let g = Goal_select.get_default_goal_selector () in let (p, status) = Pfedit.solve g None tac ?with_end_tac p in (* in case a strict subtree was completed, go back to the top of the prooftree *) diff --git a/src/tac2entries.mli b/src/tac2entries.mli index cfb58ea383..ad7624b7d0 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -13,13 +13,13 @@ open Tac2expr (** {5 Toplevel definitions} *) val register_ltac : ?local:bool -> ?mut:bool -> rec_flag -> - (Misctypes.lname * raw_tacexpr) list -> unit + (Names.lname * raw_tacexpr) list -> unit val register_type : ?local:bool -> rec_flag -> (qualid CAst.t * redef_flag * raw_quant_typedef) list -> unit val register_primitive : ?local:bool -> - Misctypes.lident -> raw_typexpr -> ml_tactic_name -> unit + Names.lident -> raw_typexpr -> ml_tactic_name -> unit val register_struct : ?local:bool -> strexpr -> unit diff --git a/src/tac2expr.mli b/src/tac2expr.mli index ddffd13a31..1f2c3ebf3b 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Loc open Names open Libnames @@ -78,7 +77,7 @@ type glb_typedef = type type_scheme = int * int glb_typexpr -type raw_quant_typedef = Misctypes.lident list * raw_typedef +type raw_quant_typedef = Names.lident list * raw_typedef type glb_quant_typedef = int * glb_typedef (** {5 Term syntax} *) @@ -159,11 +158,11 @@ type sexpr = (** {5 Toplevel statements} *) type strexpr = -| StrVal of mutable_flag * rec_flag * (Misctypes.lname * raw_tacexpr) list +| StrVal of mutable_flag * rec_flag * (Names.lname * raw_tacexpr) list (** Term definition *) | StrTyp of rec_flag * (qualid CAst.t * redef_flag * raw_quant_typedef) list (** Type definition *) -| StrPrm of Misctypes.lident * raw_typexpr * ml_tactic_name +| StrPrm of Names.lident * raw_typexpr * ml_tactic_name (** External definition *) | StrSyn of sexpr list * int option * raw_tacexpr (** Syntactic extensions *) diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 75fca938f4..86d81ef5d2 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -12,7 +12,7 @@ open CAst open CErrors open Names open Libnames -open Misctypes +open Locus open Tac2env open Tac2print open Tac2expr diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index 05b9f4141f..3f0c591e63 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -135,7 +135,7 @@ type constr_matching = constr_match_branch list CAst.t type goal_match_pattern_r = { q_goal_match_concl : constr_match_pattern; - q_goal_match_hyps : (Misctypes.lname * constr_match_pattern) list; + q_goal_match_hyps : (Names.lname * constr_match_pattern) list; } type goal_match_pattern = goal_match_pattern_r CAst.t diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml index be7c76744d..ffef2c05fd 100644 --- a/src/tac2stdlib.ml +++ b/src/tac2stdlib.ml @@ -194,10 +194,10 @@ let to_inversion_kind v = match Value.to_int v with let inversion_kind = make_to_repr to_inversion_kind let to_move_location = function -| ValInt 0 -> Misctypes.MoveFirst -| ValInt 1 -> Misctypes.MoveLast -| ValBlk (0, [|id|]) -> Misctypes.MoveAfter (Value.to_ident id) -| ValBlk (1, [|id|]) -> Misctypes.MoveBefore (Value.to_ident id) +| ValInt 0 -> Logic.MoveFirst +| ValInt 1 -> Logic.MoveLast +| ValBlk (0, [|id|]) -> Logic.MoveAfter (Value.to_ident id) +| ValBlk (1, [|id|]) -> Logic.MoveBefore (Value.to_ident id) | _ -> assert false let move_location = make_to_repr to_move_location @@ -424,7 +424,7 @@ let () = define_prim2 "tac_move" ident move_location begin fun id mv -> end let () = define_prim2 "tac_intro" (option ident) (option move_location) begin fun id mv -> - let mv = Option.default Misctypes.MoveLast mv in + let mv = Option.default Logic.MoveLast mv in Tactics.intro_move id mv end diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index 3f2385a8d6..3c464469f0 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -35,38 +35,38 @@ let delayed_of_thunk r tac env sigma = delayed_of_tactic (thaw r tac) env sigma let mk_bindings = function -| ImplicitBindings l -> Misctypes.ImplicitBindings l +| ImplicitBindings l -> Tactypes.ImplicitBindings l | ExplicitBindings l -> let l = List.map CAst.make l in - Misctypes.ExplicitBindings l -| NoBindings -> Misctypes.NoBindings + Tactypes.ExplicitBindings l +| NoBindings -> Tactypes.NoBindings let mk_with_bindings (x, b) = (x, mk_bindings b) let rec mk_intro_pattern = function -| IntroForthcoming b -> CAst.make @@ Misctypes.IntroForthcoming b -| IntroNaming ipat -> CAst.make @@ Misctypes.IntroNaming (mk_intro_pattern_naming ipat) -| IntroAction ipat -> CAst.make @@ Misctypes.IntroAction (mk_intro_pattern_action ipat) +| IntroForthcoming b -> CAst.make @@ Tactypes.IntroForthcoming b +| IntroNaming ipat -> CAst.make @@ Tactypes.IntroNaming (mk_intro_pattern_naming ipat) +| IntroAction ipat -> CAst.make @@ Tactypes.IntroAction (mk_intro_pattern_action ipat) and mk_intro_pattern_naming = function -| IntroIdentifier id -> Misctypes.IntroIdentifier id -| IntroFresh id -> Misctypes.IntroFresh id -| IntroAnonymous -> Misctypes.IntroAnonymous +| IntroIdentifier id -> Namegen.IntroIdentifier id +| IntroFresh id -> Namegen.IntroFresh id +| IntroAnonymous -> Namegen.IntroAnonymous and mk_intro_pattern_action = function -| IntroWildcard -> Misctypes.IntroWildcard -| IntroOrAndPattern ipat -> Misctypes.IntroOrAndPattern (mk_or_and_intro_pattern ipat) -| IntroInjection ipats -> Misctypes.IntroInjection (List.map mk_intro_pattern ipats) +| IntroWildcard -> Tactypes.IntroWildcard +| IntroOrAndPattern ipat -> Tactypes.IntroOrAndPattern (mk_or_and_intro_pattern ipat) +| IntroInjection ipats -> Tactypes.IntroInjection (List.map mk_intro_pattern ipats) | IntroApplyOn (c, ipat) -> let c = CAst.make @@ delayed_of_thunk Tac2ffi.constr c in - Misctypes.IntroApplyOn (c, mk_intro_pattern ipat) -| IntroRewrite b -> Misctypes.IntroRewrite b + Tactypes.IntroApplyOn (c, mk_intro_pattern ipat) +| IntroRewrite b -> Tactypes.IntroRewrite b and mk_or_and_intro_pattern = function | IntroOrPattern ipatss -> - Misctypes.IntroOrPattern (List.map (fun ipat -> List.map mk_intro_pattern ipat) ipatss) + Tactypes.IntroOrPattern (List.map (fun ipat -> List.map mk_intro_pattern ipat) ipatss) | IntroAndPattern ipats -> - Misctypes.IntroAndPattern (List.map mk_intro_pattern ipats) + Tactypes.IntroAndPattern (List.map mk_intro_pattern ipats) let mk_intro_patterns ipat = List.map mk_intro_pattern ipat @@ -77,7 +77,7 @@ let mk_occurrences f = function | OnlyOccurrences l -> Locus.OnlyOccurrences (List.map f l) let mk_occurrences_expr occ = - mk_occurrences (fun i -> Misctypes.ArgArg i) occ + mk_occurrences (fun i -> Locus.ArgArg i) occ let mk_hyp_location (id, occs, h) = ((mk_occurrences_expr occs, id), h) @@ -188,7 +188,7 @@ let forward fst tac ipat c = let assert_ = function | AssertValue (id, c) -> - let ipat = CAst.make @@ Misctypes.IntroNaming (Misctypes.IntroIdentifier id) in + let ipat = CAst.make @@ Tactypes.IntroNaming (Namegen.IntroIdentifier id) in Tactics.forward true None (Some ipat) c | AssertType (ipat, c, tac) -> let ipat = Option.map mk_intro_pattern ipat in @@ -432,8 +432,8 @@ let inversion knd arg pat ids = | Some (_, Tactics.ElimOnIdent {CAst.v=id}) -> Inv.inv_clause knd pat ids (NamedHyp id) | Some (_, Tactics.ElimOnConstr c) -> - let open Misctypes in - let anon = CAst.make @@ IntroNaming IntroAnonymous in + let open Tactypes in + let anon = CAst.make @@ IntroNaming Namegen.IntroAnonymous in Tactics.specialize c (Some anon) >>= fun () -> Tacticals.New.onLastHypId (fun id -> Inv.inv_clause knd pat ids (NamedHyp id)) end diff --git a/src/tac2types.mli b/src/tac2types.mli index a7b0ceed6e..fa31153a27 100644 --- a/src/tac2types.mli +++ b/src/tac2types.mli @@ -17,7 +17,7 @@ type advanced_flag = bool type 'a thunk = (unit, 'a) Tac2ffi.fun1 -type quantified_hypothesis = Misctypes.quantified_hypothesis = +type quantified_hypothesis = Tactypes.quantified_hypothesis = | AnonHyp of int | NamedHyp of Id.t @@ -76,7 +76,7 @@ type induction_clause = or_and_intro_pattern option * clause option -type multi = Misctypes.multi = +type multi = Equality.multi = | Precisely of int | UpTo of int | RepeatStar -- cgit v1.2.3 From 15010cea58df81a3ccfdd5a4b2a01375e34853f3 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 18 Jun 2018 16:43:42 +0200 Subject: Do not rely on the Ident vs. Qualid artificial separation. --- src/g_ltac2.ml4 | 8 ++++---- src/tac2core.ml | 7 ++++--- src/tac2entries.mli | 2 +- src/tac2qexpr.mli | 10 ++++++++-- src/tac2quote.mli | 4 ++-- 5 files changed, 19 insertions(+), 12 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index b59a5e184e..b13c036549 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -371,8 +371,8 @@ GEXTEND Gram [ [ id = Prim.ident -> CAst.make ~loc:!@loc id ] ] ; globref: - [ [ "&"; id = Prim.ident -> CAst.make ~loc:!@loc (Libnames.Ident id) - | qid = Prim.qualid -> CAst.map (fun qid -> Libnames.Qualid qid) qid + [ [ "&"; id = Prim.ident -> CAst.make ~loc:!@loc (QHypothesis id) + | qid = Prim.qualid -> CAst.map (fun qid -> QReference qid) qid ] ] ; END @@ -666,8 +666,8 @@ GEXTEND Gram ] ] ; refglobal: - [ [ "&"; id = Prim.ident -> QExpr (CAst.make ~loc:!@loc @@ Libnames.Ident id) - | qid = Prim.qualid -> QExpr (CAst.make ~loc:!@loc @@ Libnames.Qualid qid.CAst.v) + [ [ "&"; id = Prim.ident -> QExpr (CAst.make ~loc:!@loc @@ QHypothesis id) + | qid = Prim.qualid -> QExpr (CAst.make ~loc:!@loc @@ QReference qid.CAst.v) | "$"; id = Prim.ident -> QAnti (CAst.make ~loc:!@loc id) ] ] ; diff --git a/src/tac2core.ml b/src/tac2core.ml index 5f33374486..4bd294d4df 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -910,13 +910,14 @@ let () = define_ml_object Tac2quote.wit_pattern obj let () = - let intern self ist qid = match qid with - | {CAst.v=Libnames.Ident id} -> + let intern self ist ref = match ref.CAst.v with + | Tac2qexpr.QHypothesis id -> GlbVal (Globnames.VarRef id), gtypref t_reference - | {CAst.loc;v=Libnames.Qualid qid} -> + | Tac2qexpr.QReference qid -> let gr = try Nametab.locate qid with Not_found -> + let loc = ref.CAst.loc in Nametab.error_global_not_found (CAst.make ?loc qid) in GlbVal gr, gtypref t_reference diff --git a/src/tac2entries.mli b/src/tac2entries.mli index ad7624b7d0..777f3f1a43 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -74,7 +74,7 @@ val q_rewriting : rewriting Pcoq.Gram.entry val q_clause : clause Pcoq.Gram.entry val q_dispatch : dispatch Pcoq.Gram.entry val q_occurrences : occurrences Pcoq.Gram.entry -val q_reference : Libnames.reference or_anti Pcoq.Gram.entry +val q_reference : reference or_anti Pcoq.Gram.entry val q_strategy_flag : strategy_flag Pcoq.Gram.entry val q_constr_matching : constr_matching Pcoq.Gram.entry val q_goal_matching : goal_matching Pcoq.Gram.entry diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli index 3f0c591e63..400ab1a092 100644 --- a/src/tac2qexpr.mli +++ b/src/tac2qexpr.mli @@ -16,6 +16,12 @@ type 'a or_anti = | QExpr of 'a | QAnti of Id.t CAst.t +type reference_r = +| QReference of Libnames.qualid +| QHypothesis of Id.t + +type reference = reference_r CAst.t + type quantified_hypothesis = | QAnonHyp of int CAst.t | QNamedHyp of Id.t CAst.t @@ -116,8 +122,8 @@ type red_flag_r = | QFix | QCofix | QZeta -| QConst of Libnames.reference or_anti list CAst.t -| QDeltaBut of Libnames.reference or_anti list CAst.t +| QConst of reference or_anti list CAst.t +| QDeltaBut of reference or_anti list CAst.t type red_flag = red_flag_r CAst.t diff --git a/src/tac2quote.mli b/src/tac2quote.mli index cc58144901..2ce347f397 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -60,7 +60,7 @@ val of_hintdb : hintdb -> raw_tacexpr val of_move_location : move_location -> raw_tacexpr -val of_reference : Libnames.reference or_anti -> raw_tacexpr +val of_reference : reference or_anti -> raw_tacexpr val of_hyp : ?loc:Loc.t -> Id.t CAst.t -> raw_tacexpr (** id ↦ 'Control.hyp @id' *) @@ -89,7 +89,7 @@ val wit_pattern : (Constrexpr.constr_expr, Pattern.constr_pattern) Arg.tag val wit_ident : (Id.t, Id.t) Arg.tag -val wit_reference : (Libnames.reference, GlobRef.t) Arg.tag +val wit_reference : (reference, GlobRef.t) Arg.tag (** Beware, at the raw level, [Qualid [id]] has not the same meaning as [Ident id]. The first is an unqualified global reference, the second is the dynamic reference to id. *) -- cgit v1.2.3 From eba6d1ffe7a3aa775e6a4984914461364149573f Mon Sep 17 00:00:00 2001 From: Maxime Dénès Date: Mon, 18 Jun 2018 14:21:19 +0200 Subject: Adapt to Coq's PR #7797 (removal of reference). --- src/g_ltac2.ml4 | 32 ++++++++++++++++---------------- src/tac2core.ml | 3 +-- src/tac2entries.ml | 43 ++++++++++++++++++++----------------------- src/tac2entries.mli | 4 ++-- src/tac2env.ml | 4 ++-- src/tac2env.mli | 2 +- src/tac2expr.mli | 6 +++--- src/tac2intern.ml | 32 ++++++++++++++++---------------- src/tac2quote.ml | 4 ++-- src/tac2quote.mli | 3 --- 10 files changed, 63 insertions(+), 70 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index b13c036549..16e7278235 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -98,25 +98,25 @@ let inj_pattern loc c = inj_wit Tac2quote.wit_pattern loc c let inj_reference loc c = inj_wit Tac2quote.wit_reference loc c let inj_ltac1 loc e = inj_wit Tac2quote.wit_ltac1 loc e -let pattern_of_qualid ?loc id = - if Tac2env.is_constructor id.CAst.v then CAst.make ?loc @@ CPatRef (RelId id, []) +let pattern_of_qualid qid = + if Tac2env.is_constructor qid then CAst.make ?loc:qid.CAst.loc @@ CPatRef (RelId qid, []) else - let (dp, id) = Libnames.repr_qualid id.CAst.v in - if DirPath.is_empty dp then CAst.make ?loc @@ CPatVar (Name id) + let open Libnames in + if qualid_is_ident qid then CAst.make ?loc:qid.CAst.loc @@ CPatVar (Name (qualid_basename qid)) else - CErrors.user_err ?loc (Pp.str "Syntax error") + CErrors.user_err ?loc:qid.CAst.loc (Pp.str "Syntax error") GEXTEND Gram GLOBAL: tac2expr tac2type tac2def_val tac2def_typ tac2def_ext tac2def_syn tac2def_mut tac2def_run; tac2pat: [ "1" LEFTA - [ id = Prim.qualid; pl = LIST1 tac2pat LEVEL "0" -> - if Tac2env.is_constructor (id.CAst.v) then - CAst.make ~loc:!@loc @@ CPatRef (RelId id, pl) + [ qid = Prim.qualid; pl = LIST1 tac2pat LEVEL "0" -> + if Tac2env.is_constructor qid then + CAst.make ~loc:!@loc @@ CPatRef (RelId qid, pl) else CErrors.user_err ~loc:!@loc (Pp.str "Syntax error") - | id = Prim.qualid -> pattern_of_qualid ~loc:!@loc id + | qid = Prim.qualid -> pattern_of_qualid qid | "["; "]" -> CAst.make ~loc:!@loc @@ CPatRef (AbsKn (Other Tac2core.Core.c_nil), []) | p1 = tac2pat; "::"; p2 = tac2pat -> CAst.make ~loc:!@loc @@ CPatRef (AbsKn (Other Tac2core.Core.c_cons), [p1; p2]) @@ -124,7 +124,7 @@ GEXTEND Gram | "0" [ "_" -> CAst.make ~loc:!@loc @@ CPatVar Anonymous | "()" -> CAst.make ~loc:!@loc @@ CPatRef (AbsKn (Tuple 0), []) - | id = Prim.qualid -> pattern_of_qualid ~loc:!@loc id + | qid = Prim.qualid -> pattern_of_qualid qid | "("; p = atomic_tac2pat; ")" -> p ] ] ; @@ -205,11 +205,11 @@ GEXTEND Gram tactic_atom: [ [ n = Prim.integer -> CAst.make ~loc:!@loc @@ CTacAtm (AtmInt n) | s = Prim.string -> CAst.make ~loc:!@loc @@ CTacAtm (AtmStr s) - | id = Prim.qualid -> - if Tac2env.is_constructor id.CAst.v then - CAst.make ~loc:!@loc @@ CTacCst (RelId id) + | qid = Prim.qualid -> + if Tac2env.is_constructor qid then + CAst.make ~loc:!@loc @@ CTacCst (RelId qid) else - CAst.make ~loc:!@loc @@ CTacRef (RelId id) + CAst.make ~loc:!@loc @@ CTacRef (RelId qid) | "@"; id = Prim.ident -> Tac2quote.of_ident (CAst.make ~loc:!@loc id) | "&"; id = lident -> Tac2quote.of_hyp ~loc:!@loc id | "'"; c = Constr.constr -> inj_open_constr !@loc c @@ -372,7 +372,7 @@ GEXTEND Gram ; globref: [ [ "&"; id = Prim.ident -> CAst.make ~loc:!@loc (QHypothesis id) - | qid = Prim.qualid -> CAst.map (fun qid -> QReference qid) qid + | qid = Prim.qualid -> CAst.make ~loc:!@loc @@ QReference qid ] ] ; END @@ -667,7 +667,7 @@ GEXTEND Gram ; refglobal: [ [ "&"; id = Prim.ident -> QExpr (CAst.make ~loc:!@loc @@ QHypothesis id) - | qid = Prim.qualid -> QExpr (CAst.make ~loc:!@loc @@ QReference qid.CAst.v) + | qid = Prim.qualid -> QExpr (CAst.make ~loc:!@loc @@ QReference qid) | "$"; id = Prim.ident -> QAnti (CAst.make ~loc:!@loc id) ] ] ; diff --git a/src/tac2core.ml b/src/tac2core.ml index 4bd294d4df..97f25ef5ed 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -917,8 +917,7 @@ let () = let gr = try Nametab.locate qid with Not_found -> - let loc = ref.CAst.loc in - Nametab.error_global_not_found (CAst.make ?loc qid) + Nametab.error_global_not_found qid in GlbVal gr, gtypref t_reference in diff --git a/src/tac2entries.ml b/src/tac2entries.ml index c85ffb2539..8ebfeec948 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -303,7 +303,7 @@ let inline_rec_tactic tactics = let bnd = List.map map_body tactics in let pat_of_id {loc;v=id} = CAst.make ?loc @@ CPatVar (Name id) in let var_of_id {loc;v=id} = - let qid = CAst.make ?loc @@ qualid_of_ident id in + let qid = qualid_of_ident ?loc id in CAst.make ?loc @@ CTacRef (RelId qid) in let loc0 = e.loc in @@ -360,10 +360,9 @@ let register_ltac ?(local = false) ?(mut = false) isrec tactics = in List.iter iter defs -let qualid_to_ident {loc;v=qid} = - let (dp, id) = Libnames.repr_qualid qid in - if DirPath.is_empty dp then CAst.make ?loc id - else user_err ?loc (str "Identifier expected") +let qualid_to_ident qid = + if qualid_is_ident qid then CAst.make ?loc:qid.CAst.loc @@ qualid_basename qid + else user_err ?loc:qid.CAst.loc (str "Identifier expected") let register_typedef ?(local = false) isrec types = let same_name ({v=id1}, _) ({v=id2}, _) = Id.equal id1 id2 in @@ -453,21 +452,21 @@ let register_primitive ?(local = false) {loc;v=id} t ml = } in ignore (Lib.add_leaf id (inTacDef def)) -let register_open ?(local = false) {loc;v=qid} (params, def) = +let register_open ?(local = false) qid (params, def) = let kn = try Tac2env.locate_type qid with Not_found -> - user_err ?loc (str "Unbound type " ++ pr_qualid qid) + user_err ?loc:qid.CAst.loc (str "Unbound type " ++ pr_qualid qid) in let (tparams, t) = Tac2env.interp_type kn in let () = match t with | GTydOpn -> () | GTydAlg _ | GTydRec _ | GTydDef _ -> - user_err ?loc (str "Type " ++ pr_qualid qid ++ str " is not an open type") + user_err ?loc:qid.CAst.loc (str "Type " ++ pr_qualid qid ++ str " is not an open type") in let () = if not (Int.equal (List.length params) tparams) then - Tac2intern.error_nparams_mismatch ?loc (List.length params) tparams + Tac2intern.error_nparams_mismatch ?loc:qid.CAst.loc (List.length params) tparams in match def with | CTydOpn -> () @@ -492,12 +491,11 @@ let register_open ?(local = false) {loc;v=qid} (params, def) = } in Lib.add_anonymous_leaf (inTypExt def) | CTydRec _ | CTydDef _ -> - user_err ?loc (str "Extensions only accept inductive constructors") + user_err ?loc:qid.CAst.loc (str "Extensions only accept inductive constructors") let register_type ?local isrec types = match types with | [qid, true, def] -> - let {loc} = qid in - let () = if isrec then user_err ?loc (str "Extensions cannot be recursive") in + let () = if isrec then user_err ?loc:qid.CAst.loc (str "Extensions cannot be recursive") in register_open ?local qid def | _ -> let map (qid, redef, def) = @@ -709,30 +707,30 @@ let inTac2Redefinition : redefinition -> obj = subst_function = subst_redefinition; classify_function = classify_redefinition } -let register_redefinition ?(local = false) (loc, qid) e = +let register_redefinition ?(local = false) qid e = let kn = try Tac2env.locate_ltac qid - with Not_found -> user_err ?loc (str "Unknown tactic " ++ pr_qualid qid) + with Not_found -> user_err ?loc:qid.CAst.loc (str "Unknown tactic " ++ pr_qualid qid) in let kn = match kn with | TacConstant kn -> kn | TacAlias _ -> - user_err ?loc (str "Cannot redefine syntactic abbreviations") + user_err ?loc:qid.CAst.loc (str "Cannot redefine syntactic abbreviations") in let data = Tac2env.interp_global kn in let () = if not (data.Tac2env.gdata_mutable) then - user_err ?loc (str "The tactic " ++ pr_qualid qid ++ str " is not declared as mutable") + user_err ?loc:qid.CAst.loc (str "The tactic " ++ pr_qualid qid ++ str " is not declared as mutable") in let (e, t) = intern ~strict:true e in let () = if not (is_value e) then - user_err ?loc (str "Tactic definition must be a syntactical value") + user_err ?loc:qid.CAst.loc (str "Tactic definition must be a syntactical value") in let () = if not (Tac2intern.check_subtype t data.Tac2env.gdata_type) then let name = int_name () in - user_err ?loc (str "Type " ++ pr_glbtype name (snd t) ++ + user_err ?loc:qid.CAst.loc (str "Type " ++ pr_glbtype name (snd t) ++ str " is not a subtype of " ++ pr_glbtype name (snd data.Tac2env.gdata_type)) in let def = { @@ -779,7 +777,7 @@ let register_struct ?local str = match str with | StrTyp (isrec, t) -> register_type ?local isrec t | StrPrm (id, t, ml) -> register_primitive ?local id t ml | StrSyn (tok, lev, e) -> register_notation ?local tok lev e -| StrMut (qid, e) -> register_redefinition ?local CAst.(qid.loc, qid.v) e +| StrMut (qid, e) -> register_redefinition ?local qid e | StrRun e -> perform_eval e (** Toplevel exception *) @@ -831,19 +829,18 @@ end (** Printing *) -let print_ltac ref = - let {loc;v=qid} = qualid_of_reference ref in +let print_ltac qid = if Tac2env.is_constructor qid then let kn = try Tac2env.locate_constructor qid - with Not_found -> user_err ?loc (str "Unknown constructor " ++ pr_qualid qid) + with Not_found -> user_err ?loc:qid.CAst.loc (str "Unknown constructor " ++ pr_qualid qid) in let _ = Tac2env.interp_constructor kn in Feedback.msg_notice (hov 2 (str "Constructor" ++ spc () ++ str ":" ++ spc () ++ pr_qualid qid)) else let kn = try Tac2env.locate_ltac qid - with Not_found -> user_err ?loc (str "Unknown tactic " ++ pr_qualid qid) + with Not_found -> user_err ?loc:qid.CAst.loc (str "Unknown tactic " ++ pr_qualid qid) in match kn with | TacConstant kn -> diff --git a/src/tac2entries.mli b/src/tac2entries.mli index 777f3f1a43..37944981d7 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -16,7 +16,7 @@ val register_ltac : ?local:bool -> ?mut:bool -> rec_flag -> (Names.lname * raw_tacexpr) list -> unit val register_type : ?local:bool -> rec_flag -> - (qualid CAst.t * redef_flag * raw_quant_typedef) list -> unit + (qualid * redef_flag * raw_quant_typedef) list -> unit val register_primitive : ?local:bool -> Names.lident -> raw_typexpr -> ml_tactic_name -> unit @@ -41,7 +41,7 @@ val parse_scope : sexpr -> scope_rule (** {5 Inspecting} *) -val print_ltac : Libnames.reference -> unit +val print_ltac : Libnames.qualid -> unit (** {5 Eval loop} *) diff --git a/src/tac2env.ml b/src/tac2env.ml index d0f286b396..dcf7440498 100644 --- a/src/tac2env.ml +++ b/src/tac2env.ml @@ -208,10 +208,10 @@ let locate_extended_all_type qid = let tab = !nametab in KnTab.find_prefixes qid tab.tab_type -let shortest_qualid_of_type kn = +let shortest_qualid_of_type ?loc kn = let tab = !nametab in let sp = KNmap.find kn tab.tab_type_rev in - KnTab.shortest_qualid Id.Set.empty sp tab.tab_type + KnTab.shortest_qualid ?loc Id.Set.empty sp tab.tab_type let push_projection vis sp kn = let tab = !nametab in diff --git a/src/tac2env.mli b/src/tac2env.mli index 022c518143..7616579d63 100644 --- a/src/tac2env.mli +++ b/src/tac2env.mli @@ -88,7 +88,7 @@ val shortest_qualid_of_constructor : ltac_constructor -> qualid val push_type : visibility -> full_path -> type_constant -> unit val locate_type : qualid -> type_constant val locate_extended_all_type : qualid -> type_constant list -val shortest_qualid_of_type : type_constant -> qualid +val shortest_qualid_of_type : ?loc:Loc.t -> type_constant -> qualid val push_projection : visibility -> full_path -> ltac_projection -> unit val locate_projection : qualid -> ltac_projection diff --git a/src/tac2expr.mli b/src/tac2expr.mli index 1f2c3ebf3b..1069d0bfa3 100644 --- a/src/tac2expr.mli +++ b/src/tac2expr.mli @@ -26,7 +26,7 @@ type tacref = | TacAlias of ltac_alias type 'a or_relid = -| RelId of qualid CAst.t +| RelId of qualid | AbsKn of 'a (** {5 Misc} *) @@ -160,13 +160,13 @@ type sexpr = type strexpr = | StrVal of mutable_flag * rec_flag * (Names.lname * raw_tacexpr) list (** Term definition *) -| StrTyp of rec_flag * (qualid CAst.t * redef_flag * raw_quant_typedef) list +| StrTyp of rec_flag * (qualid * redef_flag * raw_quant_typedef) list (** Type definition *) | StrPrm of Names.lident * raw_typexpr * ml_tactic_name (** External definition *) | StrSyn of sexpr list * int option * raw_tacexpr (** Syntactic extensions *) -| StrMut of qualid CAst.t * raw_tacexpr +| StrMut of qualid * raw_tacexpr (** Redefinition of mutable globals *) | StrRun of raw_tacexpr (** Toplevel evaluation of an expression *) diff --git a/src/tac2intern.ml b/src/tac2intern.ml index 86d81ef5d2..f3b222df21 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -208,9 +208,9 @@ let rec intern_type env ({loc;v=t} : raw_typexpr) : UF.elt glb_typexpr = match t | CTypVar Anonymous -> GTypVar (fresh_id env) | CTypRef (rel, args) -> let (kn, nparams) = match rel with - | RelId {loc;v=qid} -> - let (dp, id) = repr_qualid qid in - if DirPath.is_empty dp && Id.Map.mem id env.env_rec then + | RelId qid -> + let id = qualid_basename qid in + if qualid_is_ident qid && Id.Map.mem id env.env_rec then let (kn, n) = Id.Map.find id env.env_rec in (Other kn, n) else @@ -230,9 +230,9 @@ let rec intern_type env ({loc;v=t} : raw_typexpr) : UF.elt glb_typexpr = match t let nargs = List.length args in let () = if not (Int.equal nparams nargs) then - let {loc;v=qid} = match rel with + let qid = match rel with | RelId lid -> lid - | AbsKn (Other kn) -> CAst.make ?loc @@ shortest_qualid_of_type kn + | AbsKn (Other kn) -> shortest_qualid_of_type ?loc kn | AbsKn (Tuple _) -> assert false in user_err ?loc (strbrk "The type constructor " ++ pr_qualid qid ++ @@ -500,14 +500,14 @@ let check_redundant_clause = function | (p, _) :: _ -> warn_redundant_clause ?loc:p.loc () let get_variable0 mem var = match var with -| RelId {loc;v=qid} -> - let (dp, id) = repr_qualid qid in - if DirPath.is_empty dp && mem id then ArgVar CAst.(make ?loc id) +| RelId qid -> + let id = qualid_basename qid in + if qualid_is_ident qid && mem id then ArgVar CAst.(make ?loc:qid.CAst.loc id) else let kn = try Tac2env.locate_ltac qid with Not_found -> - CErrors.user_err ?loc (str "Unbound value " ++ pr_qualid qid) + CErrors.user_err ?loc:qid.CAst.loc (str "Unbound value " ++ pr_qualid qid) in ArgArg kn | AbsKn kn -> ArgArg kn @@ -517,19 +517,19 @@ let get_variable env var = get_variable0 mem var let get_constructor env var = match var with -| RelId {loc;v=qid} -> +| RelId qid -> let c = try Some (Tac2env.locate_constructor qid) with Not_found -> None in begin match c with | Some knc -> Other knc | None -> - CErrors.user_err ?loc (str "Unbound constructor " ++ pr_qualid qid) + CErrors.user_err ?loc:qid.CAst.loc (str "Unbound constructor " ++ pr_qualid qid) end | AbsKn knc -> knc let get_projection var = match var with -| RelId {loc;v=qid} -> +| RelId qid -> let kn = try Tac2env.locate_projection qid with Not_found -> - user_err ?loc (pr_qualid qid ++ str " is not a projection") + user_err ?loc:qid.CAst.loc (pr_qualid qid ++ str " is not a projection") in Tac2env.interp_projection kn | AbsKn kn -> @@ -622,7 +622,7 @@ let expand_pattern avoid bnd = na, None | _ -> let id = fresh_var avoid in - let qid = RelId (CAst.make ?loc:pat.loc (qualid_of_ident id)) in + let qid = RelId (qualid_of_ident ?loc:pat.loc id) in Name id, Some qid in let avoid = ids_of_pattern avoid pat in @@ -1206,9 +1206,9 @@ let check_subtype t1 t2 = (** Globalization *) let get_projection0 var = match var with -| RelId {CAst.loc;v=qid} -> +| RelId qid -> let kn = try Tac2env.locate_projection qid with Not_found -> - user_err ?loc (pr_qualid qid ++ str " is not a projection") + user_err ?loc:qid.CAst.loc (pr_qualid qid ++ str " is not a projection") in kn | AbsKn kn -> kn diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 2a0230b779..1d742afd83 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -82,10 +82,10 @@ let inj_wit ?loc wit x = CAst.make ?loc @@ CTacExt (wit, x) let of_variable {loc;v=id} = - let qid = Libnames.qualid_of_ident id in + let qid = Libnames.qualid_of_ident ?loc id in if Tac2env.is_constructor qid then CErrors.user_err ?loc (str "Invalid identifier") - else CAst.make ?loc @@ CTacRef (RelId (CAst.make ?loc qid)) + else CAst.make ?loc @@ CTacRef (RelId qid) let of_anti f = function | QExpr x -> f x diff --git a/src/tac2quote.mli b/src/tac2quote.mli index 2ce347f397..09aa92f9ee 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -90,9 +90,6 @@ val wit_pattern : (Constrexpr.constr_expr, Pattern.constr_pattern) Arg.tag val wit_ident : (Id.t, Id.t) Arg.tag val wit_reference : (reference, GlobRef.t) Arg.tag -(** Beware, at the raw level, [Qualid [id]] has not the same meaning as - [Ident id]. The first is an unqualified global reference, the second is - the dynamic reference to id. *) val wit_constr : (Constrexpr.constr_expr, Glob_term.glob_constr) Arg.tag -- cgit v1.2.3 From fd2b270a8be038f50b57d9e76a532dcf2222fd0b Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 23 Jun 2018 14:24:19 +0200 Subject: Fix for compilation with the camlp5-parser branch. --- src/tac2core.ml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 97f25ef5ed..13265ee080 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -1127,12 +1127,12 @@ end let () = add_scope "tactic" begin function | [] -> (** Default to level 5 parsing *) - let scope = Extend.Aentryl (tac2expr, 5) in + let scope = Extend.Aentryl (tac2expr, "5") in let act tac = tac in Tac2entries.ScopeRule (scope, act) | [SexprInt {loc;v=n}] as arg -> let () = if n < 0 || n > 6 then scope_fail "tactic" arg in - let scope = Extend.Aentryl (tac2expr, n) in + let scope = Extend.Aentryl (tac2expr, string_of_int n) in let act tac = tac in Tac2entries.ScopeRule (scope, act) | arg -> scope_fail "tactic" arg -- cgit v1.2.3 From b9ad51bb8aa4fbd1bd54314797428a1a0ae19fde Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Fri, 25 May 2018 16:11:56 +0200 Subject: Adapting to move of register_constr_interp0 from Pretyping to GlobEnv. --- src/tac2core.ml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 13265ee080..d745d52c82 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -974,7 +974,7 @@ let () = let c, sigma = Pfedit.refine_by_tactic env sigma concl tac in (EConstr.of_constr c, sigma) in - Pretyping.register_constr_interp0 wit_ltac2 interp + GlobEnv.register_constr_interp0 wit_ltac2 interp let () = let interp ist env sigma concl id = @@ -984,7 +984,7 @@ let () = let sigma = Typing.check env sigma c concl in (c, sigma) in - Pretyping.register_constr_interp0 wit_ltac2_quotation interp + GlobEnv.register_constr_interp0 wit_ltac2_quotation interp let () = let pr_raw id = Genprint.PrinterBasic mt in -- cgit v1.2.3 From f37c6f514a63aa1ebfb23b3c8def0087c242ca15 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 23 Jul 2018 14:40:49 +0200 Subject: Fix deprecated warnings from Pcoq. --- src/g_ltac2.ml4 | 16 ++++++++-------- src/tac2entries.ml | 44 ++++++++++++++++++++++---------------------- src/tac2entries.mli | 42 +++++++++++++++++++++--------------------- 3 files changed, 51 insertions(+), 51 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 16e7278235..82ab53298e 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -81,14 +81,14 @@ let test_dollar_ident = end let tac2expr = Tac2entries.Pltac.tac2expr -let tac2type = Gram.entry_create "tactic:tac2type" -let tac2def_val = Gram.entry_create "tactic:tac2def_val" -let tac2def_typ = Gram.entry_create "tactic:tac2def_typ" -let tac2def_ext = Gram.entry_create "tactic:tac2def_ext" -let tac2def_syn = Gram.entry_create "tactic:tac2def_syn" -let tac2def_mut = Gram.entry_create "tactic:tac2def_mut" -let tac2def_run = Gram.entry_create "tactic:tac2def_run" -let tac2mode = Gram.entry_create "vernac:ltac2_command" +let tac2type = Entry.create "tactic:tac2type" +let tac2def_val = Entry.create "tactic:tac2def_val" +let tac2def_typ = Entry.create "tactic:tac2def_typ" +let tac2def_ext = Entry.create "tactic:tac2def_ext" +let tac2def_syn = Entry.create "tactic:tac2def_syn" +let tac2def_mut = Entry.create "tactic:tac2def_mut" +let tac2def_run = Entry.create "tactic:tac2def_run" +let tac2mode = Entry.create "vernac:ltac2_command" let ltac1_expr = Pltac.tactic_expr diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 8ebfeec948..c26fbdc478 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -22,28 +22,28 @@ open Tac2intern module Pltac = struct -let tac2expr = Pcoq.Gram.entry_create "tactic:tac2expr" - -let q_ident = Pcoq.Gram.entry_create "tactic:q_ident" -let q_bindings = Pcoq.Gram.entry_create "tactic:q_bindings" -let q_with_bindings = Pcoq.Gram.entry_create "tactic:q_with_bindings" -let q_intropattern = Pcoq.Gram.entry_create "tactic:q_intropattern" -let q_intropatterns = Pcoq.Gram.entry_create "tactic:q_intropatterns" -let q_destruction_arg = Pcoq.Gram.entry_create "tactic:q_destruction_arg" -let q_induction_clause = Pcoq.Gram.entry_create "tactic:q_induction_clause" -let q_conversion = Pcoq.Gram.entry_create "tactic:q_conversion" -let q_rewriting = Pcoq.Gram.entry_create "tactic:q_rewriting" -let q_clause = Pcoq.Gram.entry_create "tactic:q_clause" -let q_dispatch = Pcoq.Gram.entry_create "tactic:q_dispatch" -let q_occurrences = Pcoq.Gram.entry_create "tactic:q_occurrences" -let q_reference = Pcoq.Gram.entry_create "tactic:q_reference" -let q_strategy_flag = Pcoq.Gram.entry_create "tactic:q_strategy_flag" -let q_constr_matching = Pcoq.Gram.entry_create "tactic:q_constr_matching" -let q_goal_matching = Pcoq.Gram.entry_create "tactic:q_goal_matching" -let q_hintdb = Pcoq.Gram.entry_create "tactic:q_hintdb" -let q_move_location = Pcoq.Gram.entry_create "tactic:q_move_location" -let q_pose = Pcoq.Gram.entry_create "tactic:q_pose" -let q_assert = Pcoq.Gram.entry_create "tactic:q_assert" +let tac2expr = Pcoq.Entry.create "tactic:tac2expr" + +let q_ident = Pcoq.Entry.create "tactic:q_ident" +let q_bindings = Pcoq.Entry.create "tactic:q_bindings" +let q_with_bindings = Pcoq.Entry.create "tactic:q_with_bindings" +let q_intropattern = Pcoq.Entry.create "tactic:q_intropattern" +let q_intropatterns = Pcoq.Entry.create "tactic:q_intropatterns" +let q_destruction_arg = Pcoq.Entry.create "tactic:q_destruction_arg" +let q_induction_clause = Pcoq.Entry.create "tactic:q_induction_clause" +let q_conversion = Pcoq.Entry.create "tactic:q_conversion" +let q_rewriting = Pcoq.Entry.create "tactic:q_rewriting" +let q_clause = Pcoq.Entry.create "tactic:q_clause" +let q_dispatch = Pcoq.Entry.create "tactic:q_dispatch" +let q_occurrences = Pcoq.Entry.create "tactic:q_occurrences" +let q_reference = Pcoq.Entry.create "tactic:q_reference" +let q_strategy_flag = Pcoq.Entry.create "tactic:q_strategy_flag" +let q_constr_matching = Pcoq.Entry.create "tactic:q_constr_matching" +let q_goal_matching = Pcoq.Entry.create "tactic:q_goal_matching" +let q_hintdb = Pcoq.Entry.create "tactic:q_hintdb" +let q_move_location = Pcoq.Entry.create "tactic:q_move_location" +let q_pose = Pcoq.Entry.create "tactic:q_pose" +let q_assert = Pcoq.Entry.create "tactic:q_assert" end (** Tactic definition *) diff --git a/src/tac2entries.mli b/src/tac2entries.mli index 37944981d7..f97e35fec0 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -56,32 +56,32 @@ val backtrace : backtrace Exninfo.t module Pltac : sig -val tac2expr : raw_tacexpr Pcoq.Gram.entry +val tac2expr : raw_tacexpr Pcoq.Entry.t (** Quoted entries. To be used for complex notations. *) open Tac2qexpr -val q_ident : Id.t CAst.t or_anti Pcoq.Gram.entry -val q_bindings : bindings Pcoq.Gram.entry -val q_with_bindings : bindings Pcoq.Gram.entry -val q_intropattern : intro_pattern Pcoq.Gram.entry -val q_intropatterns : intro_pattern list CAst.t Pcoq.Gram.entry -val q_destruction_arg : destruction_arg Pcoq.Gram.entry -val q_induction_clause : induction_clause Pcoq.Gram.entry -val q_conversion : conversion Pcoq.Gram.entry -val q_rewriting : rewriting Pcoq.Gram.entry -val q_clause : clause Pcoq.Gram.entry -val q_dispatch : dispatch Pcoq.Gram.entry -val q_occurrences : occurrences Pcoq.Gram.entry -val q_reference : reference or_anti Pcoq.Gram.entry -val q_strategy_flag : strategy_flag Pcoq.Gram.entry -val q_constr_matching : constr_matching Pcoq.Gram.entry -val q_goal_matching : goal_matching Pcoq.Gram.entry -val q_hintdb : hintdb Pcoq.Gram.entry -val q_move_location : move_location Pcoq.Gram.entry -val q_pose : pose Pcoq.Gram.entry -val q_assert : assertion Pcoq.Gram.entry +val q_ident : Id.t CAst.t or_anti Pcoq.Entry.t +val q_bindings : bindings Pcoq.Entry.t +val q_with_bindings : bindings Pcoq.Entry.t +val q_intropattern : intro_pattern Pcoq.Entry.t +val q_intropatterns : intro_pattern list CAst.t Pcoq.Entry.t +val q_destruction_arg : destruction_arg Pcoq.Entry.t +val q_induction_clause : induction_clause Pcoq.Entry.t +val q_conversion : conversion Pcoq.Entry.t +val q_rewriting : rewriting Pcoq.Entry.t +val q_clause : clause Pcoq.Entry.t +val q_dispatch : dispatch Pcoq.Entry.t +val q_occurrences : occurrences Pcoq.Entry.t +val q_reference : reference or_anti Pcoq.Entry.t +val q_strategy_flag : strategy_flag Pcoq.Entry.t +val q_constr_matching : constr_matching Pcoq.Entry.t +val q_goal_matching : goal_matching Pcoq.Entry.t +val q_hintdb : hintdb Pcoq.Entry.t +val q_move_location : move_location Pcoq.Entry.t +val q_pose : pose Pcoq.Entry.t +val q_assert : assertion Pcoq.Entry.t end (** {5 Hooks} *) -- cgit v1.2.3 From c555a1b74ff745c7ee964c2d53463db190dc6705 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 23 Jul 2018 14:45:35 +0200 Subject: Adding environment-manipulating functions. --- _CoqProject | 1 + src/tac2core.ml | 56 +++++++++++++++++++++++++++++++++++++++++++++++++++++++ theories/Constr.v | 7 +++++++ theories/Env.v | 22 ++++++++++++++++++++++ theories/Ltac2.v | 1 + 5 files changed, 87 insertions(+) create mode 100644 theories/Env.v diff --git a/_CoqProject b/_CoqProject index 5af42197ea..15e02a6484 100644 --- a/_CoqProject +++ b/_CoqProject @@ -44,5 +44,6 @@ theories/Constr.v theories/Pattern.v theories/Fresh.v theories/Std.v +theories/Env.v theories/Notations.v theories/Ltac2.v diff --git a/src/tac2core.ml b/src/tac2core.ml index 13265ee080..8a7d2034fc 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -494,6 +494,26 @@ let () = define3 "constr_closenl" (list ident) int constr begin fun ids k c -> return (Value.of_constr ans) end +let () = define1 "constr_case" (repr_ext val_inductive) begin fun ind -> + Proofview.tclENV >>= fun env -> + try + let ans = Inductiveops.make_case_info env ind Constr.RegularStyle in + return (Value.of_ext Value.val_case ans) + with e when CErrors.noncritical e -> + throw err_notfound +end + +let () = define2 "constr_constructor" (repr_ext val_inductive) int begin fun (ind, i) k -> + Proofview.tclENV >>= fun env -> + try + let open Declarations in + let ans = Environ.lookup_mind ind env in + let _ = ans.mind_packets.(i).mind_consnames.(k) in + return (Value.of_ext val_constructor ((ind, i), (k + 1))) + with e when CErrors.noncritical e -> + throw err_notfound +end + let () = define3 "constr_in_context" ident constr closure begin fun id t c -> Proofview.Goal.goals >>= function | [gl] -> @@ -801,6 +821,42 @@ let () = define2 "fresh_fresh" (repr_ext val_free) ident begin fun avoid id -> return (Value.of_ident nid) end +(** Env *) + +let () = define1 "env_get" (list ident) begin fun ids -> + let r = match ids with + | [] -> None + | _ :: _ as ids -> + let (id, path) = List.sep_last ids in + let path = DirPath.make (List.rev path) in + let fp = Libnames.make_path path id in + try Some (Nametab.global_of_path fp) with Not_found -> None + in + return (Value.of_option Value.of_reference r) +end + +let () = define1 "env_expand" (list ident) begin fun ids -> + let r = match ids with + | [] -> [] + | _ :: _ as ids -> + let (id, path) = List.sep_last ids in + let path = DirPath.make (List.rev path) in + let qid = Libnames.make_qualid path id in + Nametab.locate_all qid + in + return (Value.of_list Value.of_reference r) +end + +let () = define1 "env_path" reference begin fun r -> + match Nametab.path_of_global r with + | fp -> + let (path, id) = Libnames.repr_path fp in + let path = DirPath.repr path in + return (Value.of_list Value.of_ident (List.rev_append path [id])) + | exception Not_found -> + throw err_notfound +end + (** ML types *) let constr_flags () = diff --git a/theories/Constr.v b/theories/Constr.v index 072c613920..d8d222730e 100644 --- a/theories/Constr.v +++ b/theories/Constr.v @@ -57,6 +57,13 @@ Ltac2 @ external closenl : ident list -> int -> constr -> constr := "ltac2" "con (** [closenl [x₁;...;xₙ] k c] abstracts over variables [x₁;...;xₙ] and replaces them with [Rel(k); ...; Rel(k+n-1)] in [c]. If two names are identical, the one of least index is kept. *) +Ltac2 @ external case : inductive -> case := "ltac2" "constr_case". +(** Generate the case information for a given inductive type. *) + +Ltac2 @ external constructor : inductive -> int -> constructor := "ltac2" "constr_constructor". +(** Generate the i-th constructor for a given inductive type. Indexing starts + at 0. Panics if there is no such constructor. *) + End Unsafe. Ltac2 @ external in_context : ident -> constr -> (unit -> unit) -> constr := "ltac2" "constr_in_context". diff --git a/theories/Env.v b/theories/Env.v new file mode 100644 index 0000000000..7e36aa7990 --- /dev/null +++ b/theories/Env.v @@ -0,0 +1,22 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Std.reference option := "ltac2" "env_get". +(** Returns the global reference corresponding to the absolute name given as + argument if it exists. *) + +Ltac2 @ external expand : ident list -> Std.reference list := "ltac2" "env_expand". +(** Returns the list of all global references whose absolute name contains + the argument list as a prefix. *) + +Ltac2 @ external path : Std.reference -> ident list := "ltac2" "env_path". +(** Returns the absolute name of the given reference. Panics if the reference + does not exist. *) diff --git a/theories/Ltac2.v b/theories/Ltac2.v index 7b2f592ac6..3fe71f4c65 100644 --- a/theories/Ltac2.v +++ b/theories/Ltac2.v @@ -18,4 +18,5 @@ Require Ltac2.Control. Require Ltac2.Fresh. Require Ltac2.Pattern. Require Ltac2.Std. +Require Ltac2.Env. Require Export Ltac2.Notations. -- cgit v1.2.3 From 6ff982c5d96cfa847f699bc25dc75553e7f718f0 Mon Sep 17 00:00:00 2001 From: Langston Barrett Date: Tue, 7 Aug 2018 11:22:32 -0700 Subject: doc/ltac2.md: add table of contents --- doc/ltac2.md | 55 +++++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 55 insertions(+) diff --git a/doc/ltac2.md b/doc/ltac2.md index 9ba227c285..e8280a033a 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -24,6 +24,61 @@ that features at least the following: This document describes the implementation of such a language. The implementation of Ltac as of Coq 8.7 will be referred to as Ltac1. +# Contents + + +**Table of Contents** + +- [Summary](#summary) +- [Contents](#contents) +- [General design](#general-design) +- [ML component](#ml-component) + - [Overview](#overview) + - [Type Syntax](#type-syntax) + - [Type declarations](#type-declarations) + - [Term Syntax](#term-syntax) + - [Ltac Definitions](#ltac-definitions) + - [Reduction](#reduction) + - [Typing](#typing) + - [Effects](#effects) + - [Standard IO](#standard-io) + - [Fatal errors](#fatal-errors) + - [Backtrack](#backtrack) + - [Goals](#goals) +- [Meta-programming](#meta-programming) + - [Overview](#overview-1) + - [Generic Syntax for Quotations](#generic-syntax-for-quotations) + - [Built-in quotations](#built-in-quotations) + - [Strict vs. non-strict mode](#strict-vs-non-strict-mode) + - [Term Antiquotations](#term-antiquotations) + - [Syntax](#syntax) + - [Semantics](#semantics) + - [Static semantics](#static-semantics) + - [Dynamic semantics](#dynamic-semantics) + - [Trivial Term Antiquotations](#trivial-term-antiquotations) + - [Match over terms](#match-over-terms) + - [Match over goals](#match-over-goals) +- [Notations](#notations) + - [Scopes](#scopes) + - [Notations](#notations-1) + - [Abbreviations](#abbreviations) +- [Evaluation](#evaluation) +- [Debug](#debug) +- [Compatibility layer with Ltac1](#compatibility-layer-with-ltac1) + - [Ltac1 from Ltac2](#ltac1-from-ltac2) + - [Ltac2 from Ltac1](#ltac2-from-ltac1) +- [Transition from Ltac1](#transition-from-ltac1) + - [Syntax changes](#syntax-changes) + - [Tactic delay](#tactic-delay) + - [Variable binding](#variable-binding) + - [In Ltac expressions](#in-ltac-expressions) + - [In quotations](#in-quotations) + - [Exception catching](#exception-catching) +- [TODO](#todo) + + + + # General design There are various alternatives to Ltac1, such that Mtac or Rtac for instance. -- cgit v1.2.3 From 057345df3a6f5e46b76b9d8a395d75ba8b0965e9 Mon Sep 17 00:00:00 2001 From: Langston Barrett Date: Tue, 7 Aug 2018 11:35:50 -0700 Subject: fix three small typos --- doc/ltac2.md | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/doc/ltac2.md b/doc/ltac2.md index 9ba227c285..38cdef025e 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -17,7 +17,7 @@ Following the need of users that start developing huge projects relying critically on Ltac, we believe that we should offer a proper modern language that features at least the following: -- at least informal, predictible semantics +- at least informal, predictable semantics - a typing system - standard programming facilities (i.e. datatypes) @@ -388,7 +388,7 @@ represent several goals, including none. Thus, there is no such thing as It is natural to do the same in Ltac2, but we must provide a way to get access to a given goal. This is the role of the `enter` primitive, that applies a -tactic to each currently focussed goal in turn. +tactic to each currently focused goal in turn. ``` val enter : (unit -> unit) -> unit @@ -634,7 +634,7 @@ bindings, so that there will be a syntax error if one of the bound variables starts with an uppercase character. The semantics of this construction is otherwise the same as the corresponding -one from Ltac1, except that it requires the goal to be focussed. +one from Ltac1, except that it requires the goal to be focused. ## Match over goals -- cgit v1.2.3 From bebd946584b33a5e263393bd88e8571cd5fa5af2 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Wed, 26 Sep 2018 11:41:05 +0200 Subject: Fix for Coq PR#8554 (term builder for tactic "change" takes an environment). --- src/tac2tactics.ml | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index 3c464469f0..25431af2ea 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -159,8 +159,7 @@ let specialize c pat = let change pat c cl = let open Tac2ffi in Proofview.Goal.enter begin fun gl -> - let env = Proofview.Goal.env gl in - let c subst sigma = + let c subst env sigma = let subst = Array.map_of_list snd (Id.Map.bindings subst) in delayed_of_tactic (Tac2ffi.app_fun1 c (array constr) constr subst) env sigma in -- cgit v1.2.3 From 118a93ccf954d02afbfa5ef2b3735ef37439b274 Mon Sep 17 00:00:00 2001 From: Maxime Dénès Date: Wed, 26 Sep 2018 16:20:07 +0200 Subject: Adapt to removal of section paths from kernel names --- src/tac2entries.ml | 4 ++-- src/tac2print.ml | 4 ++-- 2 files changed, 4 insertions(+), 4 deletions(-) diff --git a/src/tac2entries.ml b/src/tac2entries.ml index c26fbdc478..2936f143ce 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -100,8 +100,8 @@ type typdef = { } let change_kn_label kn id = - let (mp, dp, _) = KerName.repr kn in - KerName.make mp dp (Label.of_id id) + let mp = KerName.modpath kn in + KerName.make mp (Label.of_id id) let change_sp_label sp id = let (dp, _) = Libnames.repr_path sp in diff --git a/src/tac2print.ml b/src/tac2print.ml index 9c530dfc51..0fea07ee3a 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -16,8 +16,8 @@ open Tac2ffi (** Utils *) let change_kn_label kn id = - let (mp, dp, _) = KerName.repr kn in - KerName.make mp dp (Label.of_id id) + let mp = KerName.modpath kn in + KerName.make mp (Label.of_id id) let paren p = hov 2 (str "(" ++ p ++ str ")") -- cgit v1.2.3 From 0579efd93eebb18f94c90718b092ae4b68a2262d Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 2 Oct 2018 08:18:34 +0200 Subject: Fix deprecation warnings. --- src/tac2entries.ml | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/src/tac2entries.ml b/src/tac2entries.ml index c26fbdc478..95a5f954c6 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -750,14 +750,14 @@ let perform_eval e = Proof_global.give_me_the_proof () with Proof_global.NoCurrentProof -> let sigma = Evd.from_env env in - Vernacexpr.SelectAll, Proof.start sigma [] + Goal_select.SelectAll, Proof.start sigma [] in let v = match selector with - | Vernacexpr.SelectNth i -> Proofview.tclFOCUS i i v - | Vernacexpr.SelectList l -> Proofview.tclFOCUSLIST l v - | Vernacexpr.SelectId id -> Proofview.tclFOCUSID id v - | Vernacexpr.SelectAll -> v - | Vernacexpr.SelectAlreadyFocused -> assert false (** TODO **) + | Goal_select.SelectNth i -> Proofview.tclFOCUS i i v + | Goal_select.SelectList l -> Proofview.tclFOCUSLIST l v + | Goal_select.SelectId id -> Proofview.tclFOCUSID id v + | Goal_select.SelectAll -> v + | Goal_select.SelectAlreadyFocused -> assert false (** TODO **) in (** HACK: the API doesn't allow to return a value *) let ans = ref None in -- cgit v1.2.3 From e51555082231328a974ed1f7ab6a1a658df3313b Mon Sep 17 00:00:00 2001 From: Gaëtan Gilbert Date: Thu, 11 Oct 2018 14:45:34 +0200 Subject: Fix for coq/coq#8515 (command driven attributes) --- src/g_ltac2.ml4 | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 index 82ab53298e..cb42b393db 100644 --- a/src/g_ltac2.ml4 +++ b/src/g_ltac2.ml4 @@ -831,8 +831,8 @@ let classify_ltac2 = function VERNAC COMMAND FUNCTIONAL EXTEND VernacDeclareTactic2Definition | [ "Ltac2" ltac2_entry(e) ] => [ classify_ltac2 e ] -> [ - fun ~atts ~st -> let open Vernacinterp in - Tac2entries.register_struct ?local:atts.locality e; + fun ~atts ~st -> + Tac2entries.register_struct ?local:(Attributes.only_locality atts) e; st ] END -- cgit v1.2.3 From 8b9b1be48a9c83f70cbfb70f52eabc616065fa1e Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Wed, 17 Oct 2018 02:02:20 +0200 Subject: [build] Add dune file + fix warnings. This allows to drop the ltac2 folder inside the Coq dir and have it compose with the Coq build. I've fixed build warnings by the way. --- src/dune | 10 ++++++++++ src/tac2core.ml | 4 ++-- src/tac2entries.ml | 4 ++-- src/tac2ffi.ml | 2 +- src/tac2intern.ml | 2 +- src/tac2match.ml | 2 +- src/tac2print.ml | 8 ++++---- src/tac2quote.ml | 2 +- 8 files changed, 22 insertions(+), 12 deletions(-) create mode 100644 src/dune diff --git a/src/dune b/src/dune new file mode 100644 index 0000000000..b0140aa809 --- /dev/null +++ b/src/dune @@ -0,0 +1,10 @@ +(library + (name ltac2) + (public_name coq.plugins.ltac2) + (modules_without_implementation tac2expr tac2qexpr tac2types) + (libraries coq.plugins.firstorder)) + +(rule + (targets g_ltac2.ml) + (deps (:pp-file g_ltac2.ml4) ) + (action (run coqp5 -loc loc -impl %{pp-file} -o %{targets}))) diff --git a/src/tac2core.ml b/src/tac2core.ml index 6ff353ce30..8ee239f803 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -20,8 +20,8 @@ open Proofview.Notations module Value = Tac2ffi open Value -let std_core n = KerName.make2 Tac2env.std_prefix (Label.of_id (Id.of_string_soft n)) -let coq_core n = KerName.make2 Tac2env.coq_prefix (Label.of_id (Id.of_string_soft n)) +let std_core n = KerName.make Tac2env.std_prefix (Label.of_id (Id.of_string_soft n)) +let coq_core n = KerName.make Tac2env.coq_prefix (Label.of_id (Id.of_string_soft n)) module Core = struct diff --git a/src/tac2entries.ml b/src/tac2entries.ml index ec10dea777..bba4680a72 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -805,7 +805,7 @@ let pr_frame = function let () = register_handler begin function | Tac2interp.LtacError (kn, args) -> - let t_exn = KerName.make2 Tac2env.coq_prefix (Label.make "exn") in + let t_exn = KerName.make Tac2env.coq_prefix (Label.make "exn") in let v = Tac2ffi.of_open (kn, args) in let t = GTypRef (Other t_exn, []) in let c = Tac2print.pr_valexpr (Global.env ()) Evd.empty v t in @@ -897,7 +897,7 @@ let register_prim_alg name params def = let def = { typdef_local = false; typdef_expr = def } in ignore (Lib.add_leaf id (inTypDef def)) -let coq_def n = KerName.make2 Tac2env.coq_prefix (Label.make n) +let coq_def n = KerName.make Tac2env.coq_prefix (Label.make n) let def_unit = { typdef_local = false; diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index a719970a57..df1857c3e7 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -229,7 +229,7 @@ let internal_err = let coq_prefix = MPfile (DirPath.make (List.map Id.of_string ["Init"; "Ltac2"])) in - KerName.make2 coq_prefix (Label.of_id (Id.of_string "Internal")) + KerName.make coq_prefix (Label.of_id (Id.of_string "Internal")) (** FIXME: handle backtrace in Ltac2 exceptions *) let of_exn c = match fst c with diff --git a/src/tac2intern.ml b/src/tac2intern.ml index ff8fb4c0f4..fe615853ce 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -19,7 +19,7 @@ open Tac2expr (** Hardwired types and constants *) -let coq_type n = KerName.make2 Tac2env.coq_prefix (Label.make n) +let coq_type n = KerName.make Tac2env.coq_prefix (Label.make n) let t_int = coq_type "int" let t_string = coq_type "string" diff --git a/src/tac2match.ml b/src/tac2match.ml index a3140eabea..c9e549d47e 100644 --- a/src/tac2match.ml +++ b/src/tac2match.ml @@ -181,7 +181,7 @@ module PatternMatching (E:StaticEnvironment) = struct pattern_match_term pat (NamedDecl.get_type decl) >>= fun ctx -> return (id, ctx) - let hyp_match_body_and_type bodypat typepat hyps = + let _hyp_match_body_and_type bodypat typepat hyps = pick hyps >>= function | LocalDef (id,body,hyp) -> pattern_match_term bodypat body >>= fun ctx_body -> diff --git a/src/tac2print.ml b/src/tac2print.ml index 0fea07ee3a..0b20cf9f58 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -22,7 +22,7 @@ let change_kn_label kn id = let paren p = hov 2 (str "(" ++ p ++ str ")") let t_list = - KerName.make2 Tac2env.coq_prefix (Label.of_id (Id.of_string "list")) + KerName.make Tac2env.coq_prefix (Label.of_id (Id.of_string "list")) (** Type printing *) @@ -35,7 +35,7 @@ type typ_level = | T0 let t_unit = - KerName.make2 Tac2env.coq_prefix (Label.of_id (Id.of_string "unit")) + KerName.make Tac2env.coq_prefix (Label.of_id (Id.of_string "unit")) let pr_typref kn = Libnames.pr_qualid (Tac2env.shortest_qualid_of_type kn) @@ -435,7 +435,7 @@ and pr_val_list env sigma args tpe = str "[" ++ prlist_with_sep pr_semicolon pr args ++ str "]" let register_init n f = - let kn = KerName.make2 Tac2env.coq_prefix (Label.make n) in + let kn = KerName.make Tac2env.coq_prefix (Label.make n) in register_val_printer kn { val_printer = fun env sigma v _ -> f env sigma v } let () = register_init "int" begin fun _ _ n -> @@ -476,7 +476,7 @@ let () = register_init "err" begin fun _ _ e -> end let () = - let kn = KerName.make2 Tac2env.coq_prefix (Label.make "array") in + let kn = KerName.make Tac2env.coq_prefix (Label.make "array") in let val_printer env sigma v arg = match arg with | [arg] -> let (_, v) = to_block v in diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 1d742afd83..3bddfe7594 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -32,7 +32,7 @@ let control_prefix = prefix_gen "Control" let pattern_prefix = prefix_gen "Pattern" let array_prefix = prefix_gen "Array" -let kername prefix n = KerName.make2 prefix (Label.of_id (Id.of_string_soft n)) +let kername prefix n = KerName.make prefix (Label.of_id (Id.of_string_soft n)) let std_core n = kername Tac2env.std_prefix n let coq_core n = kername Tac2env.coq_prefix n let control_core n = kername control_prefix n -- cgit v1.2.3 From 0a007ed37f34c9ae30ed860e2c3f237a616c89e6 Mon Sep 17 00:00:00 2001 From: Gaëtan Gilbert Date: Tue, 30 Oct 2018 13:06:42 +0100 Subject: Adapt to coq/coq#8844 (move abstract out of tactics.ml) --- src/tac2core.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 8ee239f803..7384652216 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -781,7 +781,7 @@ let () = define1 "progress" closure begin fun f -> end let () = define2 "abstract" (option ident) closure begin fun id f -> - Tactics.tclABSTRACT id (Proofview.tclIGNORE (thaw f)) >>= fun () -> + Abstract.tclABSTRACT id (Proofview.tclIGNORE (thaw f)) >>= fun () -> return v_unit end -- cgit v1.2.3 From b06ed5af083e66ab33fbb8f77c8cce5e6b6ed2d3 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 31 Oct 2018 16:29:40 +0100 Subject: Port to coqpp. --- _CoqProject | 2 +- src/dune | 4 +- src/g_ltac2.ml4 | 868 ---------------------------------------------------- src/g_ltac2.mlg | 936 ++++++++++++++++++++++++++++++++++++++++++++++++++++++++ 4 files changed, 939 insertions(+), 871 deletions(-) delete mode 100644 src/g_ltac2.ml4 create mode 100644 src/g_ltac2.mlg diff --git a/_CoqProject b/_CoqProject index 15e02a6484..071066dd86 100644 --- a/_CoqProject +++ b/_CoqProject @@ -30,7 +30,7 @@ src/tac2tactics.ml src/tac2tactics.mli src/tac2stdlib.ml src/tac2stdlib.mli -src/g_ltac2.ml4 +src/g_ltac2.mlg src/ltac2_plugin.mlpack theories/Init.v diff --git a/src/dune b/src/dune index b0140aa809..7c911fb041 100644 --- a/src/dune +++ b/src/dune @@ -6,5 +6,5 @@ (rule (targets g_ltac2.ml) - (deps (:pp-file g_ltac2.ml4) ) - (action (run coqp5 -loc loc -impl %{pp-file} -o %{targets}))) + (deps (:mlg-file g_ltac2.mlg)) + (action (run coqpp %{mlg-file}))) diff --git a/src/g_ltac2.ml4 b/src/g_ltac2.ml4 deleted file mode 100644 index cb42b393db..0000000000 --- a/src/g_ltac2.ml4 +++ /dev/null @@ -1,868 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* Tok.t Stream.t -> int option - -let entry_of_lookahead s (lk : lookahead) = - let run strm = match lk 0 strm with None -> err () | Some _ -> () in - Gram.Entry.of_parser s run - -let (>>) (lk1 : lookahead) lk2 n strm = match lk1 n strm with -| None -> None -| Some n -> lk2 n strm - -let (<+>) (lk1 : lookahead) lk2 n strm = match lk1 n strm with -| None -> lk2 n strm -| Some n -> Some n - -let lk_kw kw n strm = match stream_nth n strm with -| KEYWORD kw' | IDENT kw' -> if String.equal kw kw' then Some (n + 1) else None -| _ -> None - -let lk_ident n strm = match stream_nth n strm with -| IDENT _ -> Some (n + 1) -| _ -> None - -let lk_int n strm = match stream_nth n strm with -| INT _ -> Some (n + 1) -| _ -> None - -let lk_ident_or_anti = lk_ident <+> (lk_kw "$" >> lk_ident) - -(* lookahead for (x:=t), (?x:=t) and (1:=t) *) -let test_lpar_idnum_coloneq = - entry_of_lookahead "test_lpar_idnum_coloneq" begin - lk_kw "(" >> (lk_ident_or_anti <+> lk_int) >> lk_kw ":=" - end - -(* lookahead for (x:t), (?x:t) *) -let test_lpar_id_colon = - entry_of_lookahead "test_lpar_id_colon" begin - lk_kw "(" >> lk_ident_or_anti >> lk_kw ":" - end - -(* Hack to recognize "(x := t)" and "($x := t)" *) -let test_lpar_id_coloneq = - entry_of_lookahead "test_lpar_id_coloneq" begin - lk_kw "(" >> lk_ident_or_anti >> lk_kw ":=" - end - -(* Hack to recognize "(x)" *) -let test_lpar_id_rpar = - entry_of_lookahead "test_lpar_id_rpar" begin - lk_kw "(" >> lk_ident >> lk_kw ")" - end - -let test_ampersand_ident = - entry_of_lookahead "test_ampersand_ident" begin - lk_kw "&" >> lk_ident - end - -let test_dollar_ident = - entry_of_lookahead "test_dollar_ident" begin - lk_kw "$" >> lk_ident - end - -let tac2expr = Tac2entries.Pltac.tac2expr -let tac2type = Entry.create "tactic:tac2type" -let tac2def_val = Entry.create "tactic:tac2def_val" -let tac2def_typ = Entry.create "tactic:tac2def_typ" -let tac2def_ext = Entry.create "tactic:tac2def_ext" -let tac2def_syn = Entry.create "tactic:tac2def_syn" -let tac2def_mut = Entry.create "tactic:tac2def_mut" -let tac2def_run = Entry.create "tactic:tac2def_run" -let tac2mode = Entry.create "vernac:ltac2_command" - -let ltac1_expr = Pltac.tactic_expr - -let inj_wit wit loc x = CAst.make ~loc @@ CTacExt (wit, x) -let inj_open_constr loc c = inj_wit Tac2quote.wit_open_constr loc c -let inj_pattern loc c = inj_wit Tac2quote.wit_pattern loc c -let inj_reference loc c = inj_wit Tac2quote.wit_reference loc c -let inj_ltac1 loc e = inj_wit Tac2quote.wit_ltac1 loc e - -let pattern_of_qualid qid = - if Tac2env.is_constructor qid then CAst.make ?loc:qid.CAst.loc @@ CPatRef (RelId qid, []) - else - let open Libnames in - if qualid_is_ident qid then CAst.make ?loc:qid.CAst.loc @@ CPatVar (Name (qualid_basename qid)) - else - CErrors.user_err ?loc:qid.CAst.loc (Pp.str "Syntax error") - -GEXTEND Gram - GLOBAL: tac2expr tac2type tac2def_val tac2def_typ tac2def_ext tac2def_syn - tac2def_mut tac2def_run; - tac2pat: - [ "1" LEFTA - [ qid = Prim.qualid; pl = LIST1 tac2pat LEVEL "0" -> - if Tac2env.is_constructor qid then - CAst.make ~loc:!@loc @@ CPatRef (RelId qid, pl) - else - CErrors.user_err ~loc:!@loc (Pp.str "Syntax error") - | qid = Prim.qualid -> pattern_of_qualid qid - | "["; "]" -> CAst.make ~loc:!@loc @@ CPatRef (AbsKn (Other Tac2core.Core.c_nil), []) - | p1 = tac2pat; "::"; p2 = tac2pat -> - CAst.make ~loc:!@loc @@ CPatRef (AbsKn (Other Tac2core.Core.c_cons), [p1; p2]) - ] - | "0" - [ "_" -> CAst.make ~loc:!@loc @@ CPatVar Anonymous - | "()" -> CAst.make ~loc:!@loc @@ CPatRef (AbsKn (Tuple 0), []) - | qid = Prim.qualid -> pattern_of_qualid qid - | "("; p = atomic_tac2pat; ")" -> p - ] ] - ; - atomic_tac2pat: - [ [ -> - CAst.make ~loc:!@loc @@ CPatRef (AbsKn (Tuple 0), []) - | p = tac2pat; ":"; t = tac2type -> - CAst.make ~loc:!@loc @@ CPatCnv (p, t) - | p = tac2pat; ","; pl = LIST0 tac2pat SEP "," -> - let pl = p :: pl in - CAst.make ~loc:!@loc @@ CPatRef (AbsKn (Tuple (List.length pl)), pl) - | p = tac2pat -> p - ] ] - ; - tac2expr: - [ "6" RIGHTA - [ e1 = SELF; ";"; e2 = SELF -> CAst.make ~loc:!@loc @@ CTacSeq (e1, e2) ] - | "5" - [ "fun"; it = LIST1 input_fun ; "=>"; body = tac2expr LEVEL "6" -> - CAst.make ~loc:!@loc @@ CTacFun (it, body) - | "let"; isrec = rec_flag; - lc = LIST1 let_clause SEP "with"; "in"; - e = tac2expr LEVEL "6" -> - CAst.make ~loc:!@loc @@ CTacLet (isrec, lc, e) - | "match"; e = tac2expr LEVEL "5"; "with"; bl = branches; "end" -> - CAst.make ~loc:!@loc @@ CTacCse (e, bl) - ] - | "4" LEFTA [ ] - | "::" RIGHTA - [ e1 = tac2expr; "::"; e2 = tac2expr -> - CAst.make ~loc:!@loc @@ CTacApp (CAst.make ~loc:!@loc @@ CTacCst (AbsKn (Other Tac2core.Core.c_cons)), [e1; e2]) - ] - | [ e0 = SELF; ","; el = LIST1 NEXT SEP "," -> - let el = e0 :: el in - CAst.make ~loc:!@loc @@ CTacApp (CAst.make ~loc:!@loc @@ CTacCst (AbsKn (Tuple (List.length el))), el) ] - | "1" LEFTA - [ e = tac2expr; el = LIST1 tac2expr LEVEL "0" -> - CAst.make ~loc:!@loc @@ CTacApp (e, el) - | e = SELF; ".("; qid = Prim.qualid; ")" -> - CAst.make ~loc:!@loc @@ CTacPrj (e, RelId qid) - | e = SELF; ".("; qid = Prim.qualid; ")"; ":="; r = tac2expr LEVEL "5" -> - CAst.make ~loc:!@loc @@ CTacSet (e, RelId qid, r) ] - | "0" - [ "("; a = SELF; ")" -> a - | "("; a = SELF; ":"; t = tac2type; ")" -> - CAst.make ~loc:!@loc @@ CTacCnv (a, t) - | "()" -> - CAst.make ~loc:!@loc @@ CTacCst (AbsKn (Tuple 0)) - | "("; ")" -> - CAst.make ~loc:!@loc @@ CTacCst (AbsKn (Tuple 0)) - | "["; a = LIST0 tac2expr LEVEL "5" SEP ";"; "]" -> - Tac2quote.of_list ~loc:!@loc (fun x -> x) a - | "{"; a = tac2rec_fieldexprs; "}" -> - CAst.make ~loc:!@loc @@ CTacRec a - | a = tactic_atom -> a ] - ] - ; - branches: - [ [ -> [] - | "|"; bl = LIST1 branch SEP "|" -> bl - | bl = LIST1 branch SEP "|" -> bl ] - ] - ; - branch: - [ [ pat = tac2pat LEVEL "1"; "=>"; e = tac2expr LEVEL "6" -> (pat, e) ] ] - ; - rec_flag: - [ [ IDENT "rec" -> true - | -> false ] ] - ; - mut_flag: - [ [ IDENT "mutable" -> true - | -> false ] ] - ; - typ_param: - [ [ "'"; id = Prim.ident -> id ] ] - ; - tactic_atom: - [ [ n = Prim.integer -> CAst.make ~loc:!@loc @@ CTacAtm (AtmInt n) - | s = Prim.string -> CAst.make ~loc:!@loc @@ CTacAtm (AtmStr s) - | qid = Prim.qualid -> - if Tac2env.is_constructor qid then - CAst.make ~loc:!@loc @@ CTacCst (RelId qid) - else - CAst.make ~loc:!@loc @@ CTacRef (RelId qid) - | "@"; id = Prim.ident -> Tac2quote.of_ident (CAst.make ~loc:!@loc id) - | "&"; id = lident -> Tac2quote.of_hyp ~loc:!@loc id - | "'"; c = Constr.constr -> inj_open_constr !@loc c - | IDENT "constr"; ":"; "("; c = Constr.lconstr; ")" -> Tac2quote.of_constr c - | IDENT "open_constr"; ":"; "("; c = Constr.lconstr; ")" -> Tac2quote.of_open_constr c - | IDENT "ident"; ":"; "("; c = lident; ")" -> Tac2quote.of_ident c - | IDENT "pattern"; ":"; "("; c = Constr.lconstr_pattern; ")" -> inj_pattern !@loc c - | IDENT "reference"; ":"; "("; c = globref; ")" -> inj_reference !@loc c - | IDENT "ltac1"; ":"; "("; qid = ltac1_expr; ")" -> inj_ltac1 !@loc qid - ] ] - ; - let_clause: - [ [ binder = let_binder; ":="; te = tac2expr -> - let (pat, fn) = binder in - let te = match fn with - | None -> te - | Some args -> CAst.make ~loc:!@loc @@ CTacFun (args, te) - in - (pat, te) - ] ] - ; - let_binder: - [ [ pats = LIST1 input_fun -> - match pats with - | [{CAst.v=CPatVar _} as pat] -> (pat, None) - | ({CAst.v=CPatVar (Name id)} as pat) :: args -> (pat, Some args) - | [pat] -> (pat, None) - | _ -> CErrors.user_err ~loc:!@loc (str "Invalid pattern") - ] ] - ; - tac2type: - [ "5" RIGHTA - [ t1 = tac2type; "->"; t2 = tac2type -> CAst.make ~loc:!@loc @@ CTypArrow (t1, t2) ] - | "2" - [ t = tac2type; "*"; tl = LIST1 tac2type LEVEL "1" SEP "*" -> - let tl = t :: tl in - CAst.make ~loc:!@loc @@ CTypRef (AbsKn (Tuple (List.length tl)), tl) ] - | "1" LEFTA - [ t = SELF; qid = Prim.qualid -> CAst.make ~loc:!@loc @@ CTypRef (RelId qid, [t]) ] - | "0" - [ "("; t = tac2type LEVEL "5"; ")" -> t - | id = typ_param -> CAst.make ~loc:!@loc @@ CTypVar (Name id) - | "_" -> CAst.make ~loc:!@loc @@ CTypVar Anonymous - | qid = Prim.qualid -> CAst.make ~loc:!@loc @@ CTypRef (RelId qid, []) - | "("; p = LIST1 tac2type LEVEL "5" SEP ","; ")"; qid = Prim.qualid -> - CAst.make ~loc:!@loc @@ CTypRef (RelId qid, p) ] - ]; - locident: - [ [ id = Prim.ident -> CAst.make ~loc:!@loc id ] ] - ; - binder: - [ [ "_" -> CAst.make ~loc:!@loc Anonymous - | l = Prim.ident -> CAst.make ~loc:!@loc (Name l) ] ] - ; - input_fun: - [ [ b = tac2pat LEVEL "0" -> b ] ] - ; - tac2def_body: - [ [ name = binder; it = LIST0 input_fun; ":="; e = tac2expr -> - let e = if List.is_empty it then e else CAst.make ~loc:!@loc @@ CTacFun (it, e) in - (name, e) - ] ] - ; - tac2def_val: - [ [ mut = mut_flag; isrec = rec_flag; l = LIST1 tac2def_body SEP "with" -> - StrVal (mut, isrec, l) - ] ] - ; - tac2def_mut: - [ [ "Set"; qid = Prim.qualid; ":="; e = tac2expr -> StrMut (qid, e) ] ] - ; - tac2def_run: - [ [ "Eval"; e = tac2expr -> StrRun e ] ] - ; - tac2typ_knd: - [ [ t = tac2type -> CTydDef (Some t) - | "["; ".."; "]" -> CTydOpn - | "["; t = tac2alg_constructors; "]" -> CTydAlg t - | "{"; t = tac2rec_fields; "}"-> CTydRec t ] ] - ; - tac2alg_constructors: - [ [ "|"; cs = LIST1 tac2alg_constructor SEP "|" -> cs - | cs = LIST0 tac2alg_constructor SEP "|" -> cs ] ] - ; - tac2alg_constructor: - [ [ c = Prim.ident -> (c, []) - | c = Prim.ident; "("; args = LIST0 tac2type SEP ","; ")"-> (c, args) ] ] - ; - tac2rec_fields: - [ [ f = tac2rec_field; ";"; l = tac2rec_fields -> f :: l - | f = tac2rec_field; ";" -> [f] - | f = tac2rec_field -> [f] - | -> [] ] ] - ; - tac2rec_field: - [ [ mut = mut_flag; id = Prim.ident; ":"; t = tac2type -> (id, mut, t) ] ] - ; - tac2rec_fieldexprs: - [ [ f = tac2rec_fieldexpr; ";"; l = tac2rec_fieldexprs -> f :: l - | f = tac2rec_fieldexpr; ";" -> [f] - | f = tac2rec_fieldexpr-> [f] - | -> [] ] ] - ; - tac2rec_fieldexpr: - [ [ qid = Prim.qualid; ":="; e = tac2expr LEVEL "1" -> RelId qid, e ] ] - ; - tac2typ_prm: - [ [ -> [] - | id = typ_param -> [CAst.make ~loc:!@loc id] - | "("; ids = LIST1 [ id = typ_param -> CAst.make ~loc:!@loc id ] SEP "," ;")" -> ids - ] ] - ; - tac2typ_def: - [ [ prm = tac2typ_prm; id = Prim.qualid; (r, e) = tac2type_body -> (id, r, (prm, e)) ] ] - ; - tac2type_body: - [ [ -> false, CTydDef None - | ":="; e = tac2typ_knd -> false, e - | "::="; e = tac2typ_knd -> true, e - ] ] - ; - tac2def_typ: - [ [ "Type"; isrec = rec_flag; l = LIST1 tac2typ_def SEP "with" -> - StrTyp (isrec, l) - ] ] - ; - tac2def_ext: - [ [ "@"; IDENT "external"; id = locident; ":"; t = tac2type LEVEL "5"; ":="; - plugin = Prim.string; name = Prim.string -> - let ml = { mltac_plugin = plugin; mltac_tactic = name } in - StrPrm (id, t, ml) - ] ] - ; - syn_node: - [ [ "_" -> CAst.make ~loc:!@loc None - | id = Prim.ident -> CAst.make ~loc:!@loc (Some id) - ] ] - ; - sexpr: - [ [ s = Prim.string -> SexprStr (CAst.make ~loc:!@loc s) - | n = Prim.integer -> SexprInt (CAst.make ~loc:!@loc n) - | id = syn_node -> SexprRec (!@loc, id, []) - | id = syn_node; "("; tok = LIST1 sexpr SEP "," ; ")" -> - SexprRec (!@loc, id, tok) - ] ] - ; - syn_level: - [ [ -> None - | ":"; n = Prim.integer -> Some n - ] ] - ; - tac2def_syn: - [ [ "Notation"; toks = LIST1 sexpr; n = syn_level; ":="; - e = tac2expr -> - StrSyn (toks, n, e) - ] ] - ; - lident: - [ [ id = Prim.ident -> CAst.make ~loc:!@loc id ] ] - ; - globref: - [ [ "&"; id = Prim.ident -> CAst.make ~loc:!@loc (QHypothesis id) - | qid = Prim.qualid -> CAst.make ~loc:!@loc @@ QReference qid - ] ] - ; -END - -(** Quotation scopes used by notations *) - -open Tac2entries.Pltac - -let loc_of_ne_list l = Loc.merge_opt (fst (List.hd l)) (fst (List.last l)) - -GEXTEND Gram - GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause - q_conversion q_rewriting q_clause q_dispatch q_occurrences q_strategy_flag - q_destruction_arg q_reference q_with_bindings q_constr_matching - q_goal_matching q_hintdb q_move_location q_pose q_assert; - anti: - [ [ "$"; id = Prim.ident -> QAnti (CAst.make ~loc:!@loc id) ] ] - ; - ident_or_anti: - [ [ id = lident -> QExpr id - | "$"; id = Prim.ident -> QAnti (CAst.make ~loc:!@loc id) - ] ] - ; - lident: - [ [ id = Prim.ident -> CAst.make ~loc:!@loc id ] ] - ; - lnatural: - [ [ n = Prim.natural -> CAst.make ~loc:!@loc n ] ] - ; - q_ident: - [ [ id = ident_or_anti -> id ] ] - ; - qhyp: - [ [ x = anti -> x - | n = lnatural -> QExpr (CAst.make ~loc:!@loc @@ QAnonHyp n) - | id = lident -> QExpr (CAst.make ~loc:!@loc @@ QNamedHyp id) - ] ] - ; - simple_binding: - [ [ "("; h = qhyp; ":="; c = Constr.lconstr; ")" -> - CAst.make ~loc:!@loc (h, c) - ] ] - ; - bindings: - [ [ test_lpar_idnum_coloneq; bl = LIST1 simple_binding -> - CAst.make ~loc:!@loc @@ QExplicitBindings bl - | bl = LIST1 Constr.constr -> - CAst.make ~loc:!@loc @@ QImplicitBindings bl - ] ] - ; - q_bindings: - [ [ bl = bindings -> bl ] ] - ; - q_with_bindings: - [ [ bl = with_bindings -> bl ] ] - ; - intropatterns: - [ [ l = LIST0 nonsimple_intropattern -> CAst.make ~loc:!@loc l ]] - ; -(* ne_intropatterns: *) -(* [ [ l = LIST1 nonsimple_intropattern -> l ]] *) -(* ; *) - or_and_intropattern: - [ [ "["; tc = LIST1 intropatterns SEP "|"; "]" -> CAst.make ~loc:!@loc @@ QIntroOrPattern tc - | "()" -> CAst.make ~loc:!@loc @@ QIntroAndPattern (CAst.make ~loc:!@loc []) - | "("; si = simple_intropattern; ")" -> CAst.make ~loc:!@loc @@ QIntroAndPattern (CAst.make ~loc:!@loc [si]) - | "("; si = simple_intropattern; ","; - tc = LIST1 simple_intropattern SEP "," ; ")" -> - CAst.make ~loc:!@loc @@ QIntroAndPattern (CAst.make ~loc:!@loc (si::tc)) - | "("; si = simple_intropattern; "&"; - tc = LIST1 simple_intropattern SEP "&" ; ")" -> - (* (A & B & C) is translated into (A,(B,C)) *) - let rec pairify = function - | ([]|[_]|[_;_]) as l -> CAst.make ~loc:!@loc l - | t::q -> - let q = - CAst.make ~loc:!@loc @@ - QIntroAction (CAst.make ~loc:!@loc @@ - QIntroOrAndPattern (CAst.make ~loc:!@loc @@ - QIntroAndPattern (pairify q))) - in - CAst.make ~loc:!@loc [t; q] - in CAst.make ~loc:!@loc @@ QIntroAndPattern (pairify (si::tc)) ] ] - ; - equality_intropattern: - [ [ "->" -> CAst.make ~loc:!@loc @@ QIntroRewrite true - | "<-" -> CAst.make ~loc:!@loc @@ QIntroRewrite false - | "[="; tc = intropatterns; "]" -> CAst.make ~loc:!@loc @@ QIntroInjection tc ] ] - ; - naming_intropattern: - [ [ LEFTQMARK; id = lident -> - CAst.make ~loc:!@loc @@ QIntroFresh (QExpr id) - | "?$"; id = lident -> - CAst.make ~loc:!@loc @@ QIntroFresh (QAnti id) - | "?" -> - CAst.make ~loc:!@loc @@ QIntroAnonymous - | id = ident_or_anti -> - CAst.make ~loc:!@loc @@ QIntroIdentifier id - ] ] - ; - nonsimple_intropattern: - [ [ l = simple_intropattern -> l - | "*" -> CAst.make ~loc:!@loc @@ QIntroForthcoming true - | "**" -> CAst.make ~loc:!@loc @@ QIntroForthcoming false ]] - ; - simple_intropattern: - [ [ pat = simple_intropattern_closed -> -(* l = LIST0 ["%"; c = operconstr LEVEL "0" -> c] -> *) - (** TODO: handle %pat *) - pat - ] ] - ; - simple_intropattern_closed: - [ [ pat = or_and_intropattern -> - CAst.make ~loc:!@loc @@ QIntroAction (CAst.make ~loc:!@loc @@ QIntroOrAndPattern pat) - | pat = equality_intropattern -> - CAst.make ~loc:!@loc @@ QIntroAction pat - | "_" -> - CAst.make ~loc:!@loc @@ QIntroAction (CAst.make ~loc:!@loc @@ QIntroWildcard) - | pat = naming_intropattern -> - CAst.make ~loc:!@loc @@ QIntroNaming pat - ] ] - ; - q_intropatterns: - [ [ ipat = intropatterns -> ipat ] ] - ; - q_intropattern: - [ [ ipat = simple_intropattern -> ipat ] ] - ; - nat_or_anti: - [ [ n = lnatural -> QExpr n - | "$"; id = Prim.ident -> QAnti (CAst.make ~loc:!@loc id) - ] ] - ; - eqn_ipat: - [ [ IDENT "eqn"; ":"; pat = naming_intropattern -> Some pat - | -> None - ] ] - ; - with_bindings: - [ [ "with"; bl = bindings -> bl | -> CAst.make ~loc:!@loc @@ QNoBindings ] ] - ; - constr_with_bindings: - [ [ c = Constr.constr; l = with_bindings -> CAst.make ~loc:!@loc @@ (c, l) ] ] - ; - destruction_arg: - [ [ n = lnatural -> CAst.make ~loc:!@loc @@ QElimOnAnonHyp n - | id = lident -> CAst.make ~loc:!@loc @@ QElimOnIdent id - | c = constr_with_bindings -> CAst.make ~loc:!@loc @@ QElimOnConstr c - ] ] - ; - q_destruction_arg: - [ [ arg = destruction_arg -> arg ] ] - ; - as_or_and_ipat: - [ [ "as"; ipat = or_and_intropattern -> Some ipat - | -> None - ] ] - ; - occs_nums: - [ [ nl = LIST1 nat_or_anti -> CAst.make ~loc:!@loc @@ QOnlyOccurrences nl - | "-"; n = nat_or_anti; nl = LIST0 nat_or_anti -> - CAst.make ~loc:!@loc @@ QAllOccurrencesBut (n::nl) - ] ] - ; - occs: - [ [ "at"; occs = occs_nums -> occs | -> CAst.make ~loc:!@loc QAllOccurrences ] ] - ; - hypident: - [ [ id = ident_or_anti -> - id,Locus.InHyp - | "("; IDENT "type"; IDENT "of"; id = ident_or_anti; ")" -> - id,Locus.InHypTypeOnly - | "("; IDENT "value"; IDENT "of"; id = ident_or_anti; ")" -> - id,Locus.InHypValueOnly - ] ] - ; - hypident_occ: - [ [ (id,l)=hypident; occs=occs -> ((occs,id),l) ] ] - ; - in_clause: - [ [ "*"; occs=occs -> - { q_onhyps = None; q_concl_occs = occs } - | "*"; "|-"; occs = concl_occ -> - { q_onhyps = None; q_concl_occs = occs } - | hl = LIST0 hypident_occ SEP ","; "|-"; occs = concl_occ -> - { q_onhyps = Some hl; q_concl_occs = occs } - | hl = LIST0 hypident_occ SEP "," -> - { q_onhyps = Some hl; q_concl_occs = CAst.make ~loc:!@loc QNoOccurrences } - ] ] - ; - clause: - [ [ "in"; cl = in_clause -> CAst.make ~loc:!@loc @@ cl - | "at"; occs = occs_nums -> - CAst.make ~loc:!@loc @@ { q_onhyps = Some []; q_concl_occs = occs } - ] ] - ; - q_clause: - [ [ cl = clause -> cl ] ] - ; - concl_occ: - [ [ "*"; occs = occs -> occs - | -> CAst.make ~loc:!@loc QNoOccurrences - ] ] - ; - induction_clause: - [ [ c = destruction_arg; pat = as_or_and_ipat; eq = eqn_ipat; - cl = OPT clause -> - CAst.make ~loc:!@loc @@ { - indcl_arg = c; - indcl_eqn = eq; - indcl_as = pat; - indcl_in = cl; - } - ] ] - ; - q_induction_clause: - [ [ cl = induction_clause -> cl ] ] - ; - conversion: - [ [ c = Constr.constr -> - CAst.make ~loc:!@loc @@ QConvert c - | c1 = Constr.constr; "with"; c2 = Constr.constr -> - CAst.make ~loc:!@loc @@ QConvertWith (c1, c2) - ] ] - ; - q_conversion: - [ [ c = conversion -> c ] ] - ; - orient: - [ [ "->" -> CAst.make ~loc:!@loc (Some true) - | "<-" -> CAst.make ~loc:!@loc (Some false) - | -> CAst.make ~loc:!@loc None - ]] - ; - rewriter: - [ [ "!"; c = constr_with_bindings -> - (CAst.make ~loc:!@loc @@ QRepeatPlus,c) - | ["?"| LEFTQMARK]; c = constr_with_bindings -> - (CAst.make ~loc:!@loc @@ QRepeatStar,c) - | n = lnatural; "!"; c = constr_with_bindings -> - (CAst.make ~loc:!@loc @@ QPrecisely n,c) - | n = lnatural; ["?" | LEFTQMARK]; c = constr_with_bindings -> - (CAst.make ~loc:!@loc @@ QUpTo n,c) - | n = lnatural; c = constr_with_bindings -> - (CAst.make ~loc:!@loc @@ QPrecisely n,c) - | c = constr_with_bindings -> - (CAst.make ~loc:!@loc @@ QPrecisely (CAst.make 1), c) - ] ] - ; - oriented_rewriter: - [ [ b = orient; (m, c) = rewriter -> - CAst.make ~loc:!@loc @@ { - rew_orient = b; - rew_repeat = m; - rew_equatn = c; - } - ] ] - ; - q_rewriting: - [ [ r = oriented_rewriter -> r ] ] - ; - tactic_then_last: - [ [ "|"; lta = LIST0 OPT tac2expr LEVEL "6" SEP "|" -> lta - | -> [] - ] ] - ; - tactic_then_gen: - [ [ ta = tac2expr; "|"; (first,last) = tactic_then_gen -> (Some ta :: first, last) - | ta = tac2expr; ".."; l = tactic_then_last -> ([], Some (Some ta, l)) - | ".."; l = tactic_then_last -> ([], Some (None, l)) - | ta = tac2expr -> ([Some ta], None) - | "|"; (first,last) = tactic_then_gen -> (None :: first, last) - | -> ([None], None) - ] ] - ; - q_dispatch: - [ [ d = tactic_then_gen -> CAst.make ~loc:!@loc d ] ] - ; - q_occurrences: - [ [ occs = occs -> occs ] ] - ; - red_flag: - [ [ IDENT "beta" -> CAst.make ~loc:!@loc @@ QBeta - | IDENT "iota" -> CAst.make ~loc:!@loc @@ QIota - | IDENT "match" -> CAst.make ~loc:!@loc @@ QMatch - | IDENT "fix" -> CAst.make ~loc:!@loc @@ QFix - | IDENT "cofix" -> CAst.make ~loc:!@loc @@ QCofix - | IDENT "zeta" -> CAst.make ~loc:!@loc @@ QZeta - | IDENT "delta"; d = delta_flag -> d - ] ] - ; - refglobal: - [ [ "&"; id = Prim.ident -> QExpr (CAst.make ~loc:!@loc @@ QHypothesis id) - | qid = Prim.qualid -> QExpr (CAst.make ~loc:!@loc @@ QReference qid) - | "$"; id = Prim.ident -> QAnti (CAst.make ~loc:!@loc id) - ] ] - ; - q_reference: - [ [ r = refglobal -> r ] ] - ; - refglobals: - [ [ gl = LIST1 refglobal -> CAst.make ~loc:!@loc gl ] ] - ; - delta_flag: - [ [ "-"; "["; idl = refglobals; "]" -> CAst.make ~loc:!@loc @@ QDeltaBut idl - | "["; idl = refglobals; "]" -> CAst.make ~loc:!@loc @@ QConst idl - | -> CAst.make ~loc:!@loc @@ QDeltaBut (CAst.make ~loc:!@loc []) - ] ] - ; - strategy_flag: - [ [ s = LIST1 red_flag -> CAst.make ~loc:!@loc s - | d = delta_flag -> - CAst.make ~loc:!@loc - [CAst.make ~loc:!@loc QBeta; CAst.make ~loc:!@loc QIota; CAst.make ~loc:!@loc QZeta; d] - ] ] - ; - q_strategy_flag: - [ [ flag = strategy_flag -> flag ] ] - ; - hintdb: - [ [ "*" -> CAst.make ~loc:!@loc @@ QHintAll - | l = LIST1 ident_or_anti -> CAst.make ~loc:!@loc @@ QHintDbs l - ] ] - ; - q_hintdb: - [ [ db = hintdb -> db ] ] - ; - match_pattern: - [ [ IDENT "context"; id = OPT Prim.ident; - "["; pat = Constr.lconstr_pattern; "]" -> CAst.make ~loc:!@loc @@ QConstrMatchContext (id, pat) - | pat = Constr.lconstr_pattern -> CAst.make ~loc:!@loc @@ QConstrMatchPattern pat ] ] - ; - match_rule: - [ [ mp = match_pattern; "=>"; tac = tac2expr -> - CAst.make ~loc:!@loc @@ (mp, tac) - ] ] - ; - match_list: - [ [ mrl = LIST1 match_rule SEP "|" -> CAst.make ~loc:!@loc @@ mrl - | "|"; mrl = LIST1 match_rule SEP "|" -> CAst.make ~loc:!@loc @@ mrl ] ] - ; - q_constr_matching: - [ [ m = match_list -> m ] ] - ; - gmatch_hyp_pattern: - [ [ na = Prim.name; ":"; pat = match_pattern -> (na, pat) ] ] - ; - gmatch_pattern: - [ [ "["; hl = LIST0 gmatch_hyp_pattern SEP ","; "|-"; p = match_pattern; "]" -> - CAst.make ~loc:!@loc @@ { - q_goal_match_concl = p; - q_goal_match_hyps = hl; - } - ] ] - ; - gmatch_rule: - [ [ mp = gmatch_pattern; "=>"; tac = tac2expr -> - CAst.make ~loc:!@loc @@ (mp, tac) - ] ] - ; - gmatch_list: - [ [ mrl = LIST1 gmatch_rule SEP "|" -> CAst.make ~loc:!@loc @@ mrl - | "|"; mrl = LIST1 gmatch_rule SEP "|" -> CAst.make ~loc:!@loc @@ mrl ] ] - ; - q_goal_matching: - [ [ m = gmatch_list -> m ] ] - ; - move_location: - [ [ "at"; IDENT "top" -> CAst.make ~loc:!@loc @@ QMoveFirst - | "at"; IDENT "bottom" -> CAst.make ~loc:!@loc @@ QMoveLast - | IDENT "after"; id = ident_or_anti -> CAst.make ~loc:!@loc @@ QMoveAfter id - | IDENT "before"; id = ident_or_anti -> CAst.make ~loc:!@loc @@ QMoveBefore id - ] ] - ; - q_move_location: - [ [ mv = move_location -> mv ] ] - ; - as_name: - [ [ -> None - | "as"; id = ident_or_anti -> Some id - ] ] - ; - pose: - [ [ test_lpar_id_coloneq; "("; id = ident_or_anti; ":="; c = Constr.lconstr; ")" -> - CAst.make ~loc:!@loc (Some id, c) - | c = Constr.constr; na = as_name -> CAst.make ~loc:!@loc (na, c) - ] ] - ; - q_pose: - [ [ p = pose -> p ] ] - ; - as_ipat: - [ [ "as"; ipat = simple_intropattern -> Some ipat - | -> None - ] ] - ; - by_tactic: - [ [ "by"; tac = tac2expr -> Some tac - | -> None - ] ] - ; - assertion: - [ [ test_lpar_id_coloneq; "("; id = ident_or_anti; ":="; c = Constr.lconstr; ")" -> - CAst.make ~loc:!@loc (QAssertValue (id, c)) - | test_lpar_id_colon; "("; id = ident_or_anti; ":"; c = Constr.lconstr; ")"; tac = by_tactic -> - let loc = !@loc in - let ipat = CAst.make ~loc @@ QIntroNaming (CAst.make ~loc @@ QIntroIdentifier id) in - CAst.make ~loc (QAssertType (Some ipat, c, tac)) - | c = Constr.constr; ipat = as_ipat; tac = by_tactic -> - CAst.make ~loc:!@loc (QAssertType (ipat, c, tac)) - ] ] - ; - q_assert: - [ [ a = assertion -> a ] ] - ; -END - -(** Extension of constr syntax *) - -let () = Hook.set Tac2entries.register_constr_quotations begin fun () -> -GEXTEND Gram - Pcoq.Constr.operconstr: LEVEL "0" - [ [ IDENT "ltac2"; ":"; "("; tac = tac2expr; ")" -> - let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in - CAst.make ~loc:!@loc (CHole (None, Namegen.IntroAnonymous, Some arg)) - | test_ampersand_ident; "&"; id = Prim.ident -> - let tac = Tac2quote.of_exact_hyp ~loc:!@loc (CAst.make ~loc:!@loc id) in - let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in - CAst.make ~loc:!@loc (CHole (None, Namegen.IntroAnonymous, Some arg)) - | test_dollar_ident; "$"; id = Prim.ident -> - let id = Loc.tag ~loc:!@loc id in - let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2_quotation) id in - CAst.make ~loc:!@loc (CHole (None, Namegen.IntroAnonymous, Some arg)) - ] ] - ; -END -end - -let pr_ltac2entry _ = mt () (** FIXME *) -let pr_ltac2expr _ = mt () (** FIXME *) - -VERNAC ARGUMENT EXTEND ltac2_entry -PRINTED BY pr_ltac2entry -| [ tac2def_val(v) ] -> [ v ] -| [ tac2def_typ(t) ] -> [ t ] -| [ tac2def_ext(e) ] -> [ e ] -| [ tac2def_syn(e) ] -> [ e ] -| [ tac2def_mut(e) ] -> [ e ] -| [ tac2def_run(e) ] -> [ e ] -END - -let classify_ltac2 = function -| StrSyn _ -> Vernacexpr.VtUnknown, Vernacexpr.VtNow -| StrMut _ | StrVal _ | StrPrm _ | StrTyp _ | StrRun _ -> Vernac_classifier.classify_as_sideeff - -VERNAC COMMAND FUNCTIONAL EXTEND VernacDeclareTactic2Definition -| [ "Ltac2" ltac2_entry(e) ] => [ classify_ltac2 e ] -> [ - fun ~atts ~st -> - Tac2entries.register_struct ?local:(Attributes.only_locality atts) e; - st - ] -END - -let _ = - let mode = { - Proof_global.name = "Ltac2"; - set = (fun () -> Pvernac.set_command_entry tac2mode); - reset = (fun () -> Pvernac.(set_command_entry Vernac_.noedit_mode)); - } in - Proof_global.register_proof_mode mode - -VERNAC ARGUMENT EXTEND ltac2_expr -PRINTED BY pr_ltac2expr -| [ tac2expr(e) ] -> [ e ] -END - -open G_ltac -open Vernac_classifier - -VERNAC tac2mode EXTEND VernacLtac2 -| [ - ltac2_expr(t) ltac_use_default(default) ] => - [ classify_as_proofstep ] -> [ -(* let g = Option.default (Proof_global.get_default_goal_selector ()) g in *) - Tac2entries.call ~default t - ] -END - -open Stdarg - -VERNAC COMMAND EXTEND Ltac2Print CLASSIFIED AS SIDEFF -| [ "Print" "Ltac2" reference(tac) ] -> [ Tac2entries.print_ltac tac ] -END diff --git a/src/g_ltac2.mlg b/src/g_ltac2.mlg new file mode 100644 index 0000000000..5aad77596d --- /dev/null +++ b/src/g_ltac2.mlg @@ -0,0 +1,936 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Tok.t Stream.t -> int option + +let entry_of_lookahead s (lk : lookahead) = + let run strm = match lk 0 strm with None -> err () | Some _ -> () in + Gram.Entry.of_parser s run + +let (>>) (lk1 : lookahead) lk2 n strm = match lk1 n strm with +| None -> None +| Some n -> lk2 n strm + +let (<+>) (lk1 : lookahead) lk2 n strm = match lk1 n strm with +| None -> lk2 n strm +| Some n -> Some n + +let lk_kw kw n strm = match stream_nth n strm with +| KEYWORD kw' | IDENT kw' -> if String.equal kw kw' then Some (n + 1) else None +| _ -> None + +let lk_ident n strm = match stream_nth n strm with +| IDENT _ -> Some (n + 1) +| _ -> None + +let lk_int n strm = match stream_nth n strm with +| INT _ -> Some (n + 1) +| _ -> None + +let lk_ident_or_anti = lk_ident <+> (lk_kw "$" >> lk_ident) + +(* lookahead for (x:=t), (?x:=t) and (1:=t) *) +let test_lpar_idnum_coloneq = + entry_of_lookahead "test_lpar_idnum_coloneq" begin + lk_kw "(" >> (lk_ident_or_anti <+> lk_int) >> lk_kw ":=" + end + +(* lookahead for (x:t), (?x:t) *) +let test_lpar_id_colon = + entry_of_lookahead "test_lpar_id_colon" begin + lk_kw "(" >> lk_ident_or_anti >> lk_kw ":" + end + +(* Hack to recognize "(x := t)" and "($x := t)" *) +let test_lpar_id_coloneq = + entry_of_lookahead "test_lpar_id_coloneq" begin + lk_kw "(" >> lk_ident_or_anti >> lk_kw ":=" + end + +(* Hack to recognize "(x)" *) +let test_lpar_id_rpar = + entry_of_lookahead "test_lpar_id_rpar" begin + lk_kw "(" >> lk_ident >> lk_kw ")" + end + +let test_ampersand_ident = + entry_of_lookahead "test_ampersand_ident" begin + lk_kw "&" >> lk_ident + end + +let test_dollar_ident = + entry_of_lookahead "test_dollar_ident" begin + lk_kw "$" >> lk_ident + end + +let tac2expr = Tac2entries.Pltac.tac2expr +let tac2type = Entry.create "tactic:tac2type" +let tac2def_val = Entry.create "tactic:tac2def_val" +let tac2def_typ = Entry.create "tactic:tac2def_typ" +let tac2def_ext = Entry.create "tactic:tac2def_ext" +let tac2def_syn = Entry.create "tactic:tac2def_syn" +let tac2def_mut = Entry.create "tactic:tac2def_mut" +let tac2def_run = Entry.create "tactic:tac2def_run" +let tac2mode = Entry.create "vernac:ltac2_command" + +let ltac1_expr = Pltac.tactic_expr + +let inj_wit wit loc x = CAst.make ~loc @@ CTacExt (wit, x) +let inj_open_constr loc c = inj_wit Tac2quote.wit_open_constr loc c +let inj_pattern loc c = inj_wit Tac2quote.wit_pattern loc c +let inj_reference loc c = inj_wit Tac2quote.wit_reference loc c +let inj_ltac1 loc e = inj_wit Tac2quote.wit_ltac1 loc e + +let pattern_of_qualid qid = + if Tac2env.is_constructor qid then CAst.make ?loc:qid.CAst.loc @@ CPatRef (RelId qid, []) + else + let open Libnames in + if qualid_is_ident qid then CAst.make ?loc:qid.CAst.loc @@ CPatVar (Name (qualid_basename qid)) + else + CErrors.user_err ?loc:qid.CAst.loc (Pp.str "Syntax error") + +} + +GRAMMAR EXTEND Gram + GLOBAL: tac2expr tac2type tac2def_val tac2def_typ tac2def_ext tac2def_syn + tac2def_mut tac2def_run; + tac2pat: + [ "1" LEFTA + [ qid = Prim.qualid; pl = LIST1 tac2pat LEVEL "0" -> { + if Tac2env.is_constructor qid then + CAst.make ~loc @@ CPatRef (RelId qid, pl) + else + CErrors.user_err ~loc (Pp.str "Syntax error") } + | qid = Prim.qualid -> { pattern_of_qualid qid } + | "["; "]" -> { CAst.make ~loc @@ CPatRef (AbsKn (Other Tac2core.Core.c_nil), []) } + | p1 = tac2pat; "::"; p2 = tac2pat -> + { CAst.make ~loc @@ CPatRef (AbsKn (Other Tac2core.Core.c_cons), [p1; p2])} + ] + | "0" + [ "_" -> { CAst.make ~loc @@ CPatVar Anonymous } + | "()" -> { CAst.make ~loc @@ CPatRef (AbsKn (Tuple 0), []) } + | qid = Prim.qualid -> { pattern_of_qualid qid } + | "("; p = atomic_tac2pat; ")" -> { p } + ] ] + ; + atomic_tac2pat: + [ [ -> + { CAst.make ~loc @@ CPatRef (AbsKn (Tuple 0), []) } + | p = tac2pat; ":"; t = tac2type -> + { CAst.make ~loc @@ CPatCnv (p, t) } + | p = tac2pat; ","; pl = LIST0 tac2pat SEP "," -> + { let pl = p :: pl in + CAst.make ~loc @@ CPatRef (AbsKn (Tuple (List.length pl)), pl) } + | p = tac2pat -> { p } + ] ] + ; + tac2expr: + [ "6" RIGHTA + [ e1 = SELF; ";"; e2 = SELF -> { CAst.make ~loc @@ CTacSeq (e1, e2) } ] + | "5" + [ "fun"; it = LIST1 input_fun ; "=>"; body = tac2expr LEVEL "6" -> + { CAst.make ~loc @@ CTacFun (it, body) } + | "let"; isrec = rec_flag; + lc = LIST1 let_clause SEP "with"; "in"; + e = tac2expr LEVEL "6" -> + { CAst.make ~loc @@ CTacLet (isrec, lc, e) } + | "match"; e = tac2expr LEVEL "5"; "with"; bl = branches; "end" -> + { CAst.make ~loc @@ CTacCse (e, bl) } + ] + | "4" LEFTA [ ] + | "::" RIGHTA + [ e1 = tac2expr; "::"; e2 = tac2expr -> + { CAst.make ~loc @@ CTacApp (CAst.make ~loc @@ CTacCst (AbsKn (Other Tac2core.Core.c_cons)), [e1; e2]) } + ] + | [ e0 = SELF; ","; el = LIST1 NEXT SEP "," -> + { let el = e0 :: el in + CAst.make ~loc @@ CTacApp (CAst.make ~loc @@ CTacCst (AbsKn (Tuple (List.length el))), el) } ] + | "1" LEFTA + [ e = tac2expr; el = LIST1 tac2expr LEVEL "0" -> + { CAst.make ~loc @@ CTacApp (e, el) } + | e = SELF; ".("; qid = Prim.qualid; ")" -> + { CAst.make ~loc @@ CTacPrj (e, RelId qid) } + | e = SELF; ".("; qid = Prim.qualid; ")"; ":="; r = tac2expr LEVEL "5" -> + { CAst.make ~loc @@ CTacSet (e, RelId qid, r) } ] + | "0" + [ "("; a = SELF; ")" -> { a } + | "("; a = SELF; ":"; t = tac2type; ")" -> + { CAst.make ~loc @@ CTacCnv (a, t) } + | "()" -> + { CAst.make ~loc @@ CTacCst (AbsKn (Tuple 0)) } + | "("; ")" -> + { CAst.make ~loc @@ CTacCst (AbsKn (Tuple 0)) } + | "["; a = LIST0 tac2expr LEVEL "5" SEP ";"; "]" -> + { Tac2quote.of_list ~loc (fun x -> x) a } + | "{"; a = tac2rec_fieldexprs; "}" -> + { CAst.make ~loc @@ CTacRec a } + | a = tactic_atom -> { a } ] + ] + ; + branches: + [ [ -> { [] } + | "|"; bl = LIST1 branch SEP "|" -> { bl } + | bl = LIST1 branch SEP "|" -> { bl } ] + ] + ; + branch: + [ [ pat = tac2pat LEVEL "1"; "=>"; e = tac2expr LEVEL "6" -> { (pat, e) } ] ] + ; + rec_flag: + [ [ IDENT "rec" -> { true } + | -> { false } ] ] + ; + mut_flag: + [ [ IDENT "mutable" -> { true } + | -> { false } ] ] + ; + typ_param: + [ [ "'"; id = Prim.ident -> { id } ] ] + ; + tactic_atom: + [ [ n = Prim.integer -> { CAst.make ~loc @@ CTacAtm (AtmInt n) } + | s = Prim.string -> { CAst.make ~loc @@ CTacAtm (AtmStr s) } + | qid = Prim.qualid -> + { if Tac2env.is_constructor qid then + CAst.make ~loc @@ CTacCst (RelId qid) + else + CAst.make ~loc @@ CTacRef (RelId qid) } + | "@"; id = Prim.ident -> { Tac2quote.of_ident (CAst.make ~loc id) } + | "&"; id = lident -> { Tac2quote.of_hyp ~loc id } + | "'"; c = Constr.constr -> { inj_open_constr loc c } + | IDENT "constr"; ":"; "("; c = Constr.lconstr; ")" -> { Tac2quote.of_constr c } + | IDENT "open_constr"; ":"; "("; c = Constr.lconstr; ")" -> { Tac2quote.of_open_constr c } + | IDENT "ident"; ":"; "("; c = lident; ")" -> { Tac2quote.of_ident c } + | IDENT "pattern"; ":"; "("; c = Constr.lconstr_pattern; ")" -> { inj_pattern loc c } + | IDENT "reference"; ":"; "("; c = globref; ")" -> { inj_reference loc c } + | IDENT "ltac1"; ":"; "("; qid = ltac1_expr; ")" -> { inj_ltac1 loc qid } + ] ] + ; + let_clause: + [ [ binder = let_binder; ":="; te = tac2expr -> + { let (pat, fn) = binder in + let te = match fn with + | None -> te + | Some args -> CAst.make ~loc @@ CTacFun (args, te) + in + (pat, te) } + ] ] + ; + let_binder: + [ [ pats = LIST1 input_fun -> + { match pats with + | [{CAst.v=CPatVar _} as pat] -> (pat, None) + | ({CAst.v=CPatVar (Name id)} as pat) :: args -> (pat, Some args) + | [pat] -> (pat, None) + | _ -> CErrors.user_err ~loc (str "Invalid pattern") } + ] ] + ; + tac2type: + [ "5" RIGHTA + [ t1 = tac2type; "->"; t2 = tac2type -> { CAst.make ~loc @@ CTypArrow (t1, t2) } ] + | "2" + [ t = tac2type; "*"; tl = LIST1 tac2type LEVEL "1" SEP "*" -> + { let tl = t :: tl in + CAst.make ~loc @@ CTypRef (AbsKn (Tuple (List.length tl)), tl) } ] + | "1" LEFTA + [ t = SELF; qid = Prim.qualid -> { CAst.make ~loc @@ CTypRef (RelId qid, [t]) } ] + | "0" + [ "("; t = tac2type LEVEL "5"; ")" -> { t } + | id = typ_param -> { CAst.make ~loc @@ CTypVar (Name id) } + | "_" -> { CAst.make ~loc @@ CTypVar Anonymous } + | qid = Prim.qualid -> { CAst.make ~loc @@ CTypRef (RelId qid, []) } + | "("; p = LIST1 tac2type LEVEL "5" SEP ","; ")"; qid = Prim.qualid -> + { CAst.make ~loc @@ CTypRef (RelId qid, p) } ] + ]; + locident: + [ [ id = Prim.ident -> { CAst.make ~loc id } ] ] + ; + binder: + [ [ "_" -> { CAst.make ~loc Anonymous } + | l = Prim.ident -> { CAst.make ~loc (Name l) } ] ] + ; + input_fun: + [ [ b = tac2pat LEVEL "0" -> { b } ] ] + ; + tac2def_body: + [ [ name = binder; it = LIST0 input_fun; ":="; e = tac2expr -> + { let e = if List.is_empty it then e else CAst.make ~loc @@ CTacFun (it, e) in + (name, e) } + ] ] + ; + tac2def_val: + [ [ mut = mut_flag; isrec = rec_flag; l = LIST1 tac2def_body SEP "with" -> + { StrVal (mut, isrec, l) } + ] ] + ; + tac2def_mut: + [ [ "Set"; qid = Prim.qualid; ":="; e = tac2expr -> { StrMut (qid, e) } ] ] + ; + tac2def_run: + [ [ "Eval"; e = tac2expr -> { StrRun e } ] ] + ; + tac2typ_knd: + [ [ t = tac2type -> { CTydDef (Some t) } + | "["; ".."; "]" -> { CTydOpn } + | "["; t = tac2alg_constructors; "]" -> { CTydAlg t } + | "{"; t = tac2rec_fields; "}"-> { CTydRec t } ] ] + ; + tac2alg_constructors: + [ [ "|"; cs = LIST1 tac2alg_constructor SEP "|" -> { cs } + | cs = LIST0 tac2alg_constructor SEP "|" -> { cs } ] ] + ; + tac2alg_constructor: + [ [ c = Prim.ident -> { (c, []) } + | c = Prim.ident; "("; args = LIST0 tac2type SEP ","; ")"-> { (c, args) } ] ] + ; + tac2rec_fields: + [ [ f = tac2rec_field; ";"; l = tac2rec_fields -> { f :: l } + | f = tac2rec_field; ";" -> { [f] } + | f = tac2rec_field -> { [f] } + | -> { [] } ] ] + ; + tac2rec_field: + [ [ mut = mut_flag; id = Prim.ident; ":"; t = tac2type -> { (id, mut, t) } ] ] + ; + tac2rec_fieldexprs: + [ [ f = tac2rec_fieldexpr; ";"; l = tac2rec_fieldexprs -> { f :: l } + | f = tac2rec_fieldexpr; ";" -> { [f] } + | f = tac2rec_fieldexpr-> { [f] } + | -> { [] } ] ] + ; + tac2rec_fieldexpr: + [ [ qid = Prim.qualid; ":="; e = tac2expr LEVEL "1" -> { RelId qid, e } ] ] + ; + tac2typ_prm: + [ [ -> { [] } + | id = typ_param -> { [CAst.make ~loc id] } + | "("; ids = LIST1 [ id = typ_param -> { CAst.make ~loc id } ] SEP "," ;")" -> { ids } + ] ] + ; + tac2typ_def: + [ [ prm = tac2typ_prm; id = Prim.qualid; b = tac2type_body -> { let (r, e) = b in (id, r, (prm, e)) } ] ] + ; + tac2type_body: + [ [ -> { false, CTydDef None } + | ":="; e = tac2typ_knd -> { false, e } + | "::="; e = tac2typ_knd -> { true, e } + ] ] + ; + tac2def_typ: + [ [ "Type"; isrec = rec_flag; l = LIST1 tac2typ_def SEP "with" -> + { StrTyp (isrec, l) } + ] ] + ; + tac2def_ext: + [ [ "@"; IDENT "external"; id = locident; ":"; t = tac2type LEVEL "5"; ":="; + plugin = Prim.string; name = Prim.string -> + { let ml = { mltac_plugin = plugin; mltac_tactic = name } in + StrPrm (id, t, ml) } + ] ] + ; + syn_node: + [ [ "_" -> { CAst.make ~loc None } + | id = Prim.ident -> { CAst.make ~loc (Some id) } + ] ] + ; + sexpr: + [ [ s = Prim.string -> { SexprStr (CAst.make ~loc s) } + | n = Prim.integer -> { SexprInt (CAst.make ~loc n) } + | id = syn_node -> { SexprRec (loc, id, []) } + | id = syn_node; "("; tok = LIST1 sexpr SEP "," ; ")" -> + { SexprRec (loc, id, tok) } + ] ] + ; + syn_level: + [ [ -> { None } + | ":"; n = Prim.integer -> { Some n } + ] ] + ; + tac2def_syn: + [ [ "Notation"; toks = LIST1 sexpr; n = syn_level; ":="; + e = tac2expr -> + { StrSyn (toks, n, e) } + ] ] + ; + lident: + [ [ id = Prim.ident -> { CAst.make ~loc id } ] ] + ; + globref: + [ [ "&"; id = Prim.ident -> { CAst.make ~loc (QHypothesis id) } + | qid = Prim.qualid -> { CAst.make ~loc @@ QReference qid } + ] ] + ; +END + +(** Quotation scopes used by notations *) + +{ + +open Tac2entries.Pltac + +let loc_of_ne_list l = Loc.merge_opt (fst (List.hd l)) (fst (List.last l)) + +} + +GRAMMAR EXTEND Gram + GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause + q_conversion q_rewriting q_clause q_dispatch q_occurrences q_strategy_flag + q_destruction_arg q_reference q_with_bindings q_constr_matching + q_goal_matching q_hintdb q_move_location q_pose q_assert; + anti: + [ [ "$"; id = Prim.ident -> { QAnti (CAst.make ~loc id) } ] ] + ; + ident_or_anti: + [ [ id = lident -> { QExpr id } + | "$"; id = Prim.ident -> { QAnti (CAst.make ~loc id) } + ] ] + ; + lident: + [ [ id = Prim.ident -> { CAst.make ~loc id } ] ] + ; + lnatural: + [ [ n = Prim.natural -> { CAst.make ~loc n } ] ] + ; + q_ident: + [ [ id = ident_or_anti -> { id } ] ] + ; + qhyp: + [ [ x = anti -> { x } + | n = lnatural -> { QExpr (CAst.make ~loc @@ QAnonHyp n) } + | id = lident -> { QExpr (CAst.make ~loc @@ QNamedHyp id) } + ] ] + ; + simple_binding: + [ [ "("; h = qhyp; ":="; c = Constr.lconstr; ")" -> + { CAst.make ~loc (h, c) } + ] ] + ; + bindings: + [ [ test_lpar_idnum_coloneq; bl = LIST1 simple_binding -> + { CAst.make ~loc @@ QExplicitBindings bl } + | bl = LIST1 Constr.constr -> + { CAst.make ~loc @@ QImplicitBindings bl } + ] ] + ; + q_bindings: + [ [ bl = bindings -> { bl } ] ] + ; + q_with_bindings: + [ [ bl = with_bindings -> { bl } ] ] + ; + intropatterns: + [ [ l = LIST0 nonsimple_intropattern -> { CAst.make ~loc l } ] ] + ; +(* ne_intropatterns: *) +(* [ [ l = LIST1 nonsimple_intropattern -> l ]] *) +(* ; *) + or_and_intropattern: + [ [ "["; tc = LIST1 intropatterns SEP "|"; "]" -> { CAst.make ~loc @@ QIntroOrPattern tc } + | "()" -> { CAst.make ~loc @@ QIntroAndPattern (CAst.make ~loc []) } + | "("; si = simple_intropattern; ")" -> { CAst.make ~loc @@ QIntroAndPattern (CAst.make ~loc [si]) } + | "("; si = simple_intropattern; ","; + tc = LIST1 simple_intropattern SEP "," ; ")" -> + { CAst.make ~loc @@ QIntroAndPattern (CAst.make ~loc (si::tc)) } + | "("; si = simple_intropattern; "&"; + tc = LIST1 simple_intropattern SEP "&" ; ")" -> + (* (A & B & C) is translated into (A,(B,C)) *) + { let rec pairify = function + | ([]|[_]|[_;_]) as l -> CAst.make ~loc l + | t::q -> + let q = + CAst.make ~loc @@ + QIntroAction (CAst.make ~loc @@ + QIntroOrAndPattern (CAst.make ~loc @@ + QIntroAndPattern (pairify q))) + in + CAst.make ~loc [t; q] + in CAst.make ~loc @@ QIntroAndPattern (pairify (si::tc)) } ] ] + ; + equality_intropattern: + [ [ "->" -> { CAst.make ~loc @@ QIntroRewrite true } + | "<-" -> { CAst.make ~loc @@ QIntroRewrite false } + | "[="; tc = intropatterns; "]" -> { CAst.make ~loc @@ QIntroInjection tc } ] ] + ; + naming_intropattern: + [ [ LEFTQMARK; id = lident -> + { CAst.make ~loc @@ QIntroFresh (QExpr id) } + | "?$"; id = lident -> + { CAst.make ~loc @@ QIntroFresh (QAnti id) } + | "?" -> + { CAst.make ~loc @@ QIntroAnonymous } + | id = ident_or_anti -> + { CAst.make ~loc @@ QIntroIdentifier id } + ] ] + ; + nonsimple_intropattern: + [ [ l = simple_intropattern -> { l } + | "*" -> { CAst.make ~loc @@ QIntroForthcoming true } + | "**" -> { CAst.make ~loc @@ QIntroForthcoming false } ] ] + ; + simple_intropattern: + [ [ pat = simple_intropattern_closed -> +(* l = LIST0 ["%"; c = operconstr LEVEL "0" -> c] -> *) + (** TODO: handle %pat *) + { pat } + ] ] + ; + simple_intropattern_closed: + [ [ pat = or_and_intropattern -> + { CAst.make ~loc @@ QIntroAction (CAst.make ~loc @@ QIntroOrAndPattern pat) } + | pat = equality_intropattern -> + { CAst.make ~loc @@ QIntroAction pat } + | "_" -> + { CAst.make ~loc @@ QIntroAction (CAst.make ~loc @@ QIntroWildcard) } + | pat = naming_intropattern -> + { CAst.make ~loc @@ QIntroNaming pat } + ] ] + ; + q_intropatterns: + [ [ ipat = intropatterns -> { ipat } ] ] + ; + q_intropattern: + [ [ ipat = simple_intropattern -> { ipat } ] ] + ; + nat_or_anti: + [ [ n = lnatural -> { QExpr n } + | "$"; id = Prim.ident -> { QAnti (CAst.make ~loc id) } + ] ] + ; + eqn_ipat: + [ [ IDENT "eqn"; ":"; pat = naming_intropattern -> { Some pat } + | -> { None } + ] ] + ; + with_bindings: + [ [ "with"; bl = bindings -> { bl } | -> { CAst.make ~loc @@ QNoBindings } ] ] + ; + constr_with_bindings: + [ [ c = Constr.constr; l = with_bindings -> { CAst.make ~loc @@ (c, l) } ] ] + ; + destruction_arg: + [ [ n = lnatural -> { CAst.make ~loc @@ QElimOnAnonHyp n } + | id = lident -> { CAst.make ~loc @@ QElimOnIdent id } + | c = constr_with_bindings -> { CAst.make ~loc @@ QElimOnConstr c } + ] ] + ; + q_destruction_arg: + [ [ arg = destruction_arg -> { arg } ] ] + ; + as_or_and_ipat: + [ [ "as"; ipat = or_and_intropattern -> { Some ipat } + | -> { None } + ] ] + ; + occs_nums: + [ [ nl = LIST1 nat_or_anti -> { CAst.make ~loc @@ QOnlyOccurrences nl } + | "-"; n = nat_or_anti; nl = LIST0 nat_or_anti -> + { CAst.make ~loc @@ QAllOccurrencesBut (n::nl) } + ] ] + ; + occs: + [ [ "at"; occs = occs_nums -> { occs } | -> { CAst.make ~loc QAllOccurrences } ] ] + ; + hypident: + [ [ id = ident_or_anti -> + { id,Locus.InHyp } + | "("; IDENT "type"; IDENT "of"; id = ident_or_anti; ")" -> + { id,Locus.InHypTypeOnly } + | "("; IDENT "value"; IDENT "of"; id = ident_or_anti; ")" -> + { id,Locus.InHypValueOnly } + ] ] + ; + hypident_occ: + [ [ h=hypident; occs=occs -> { let (id,l) = h in ((occs,id),l) } ] ] + ; + in_clause: + [ [ "*"; occs=occs -> + { { q_onhyps = None; q_concl_occs = occs } } + | "*"; "|-"; occs = concl_occ -> + { { q_onhyps = None; q_concl_occs = occs } } + | hl = LIST0 hypident_occ SEP ","; "|-"; occs = concl_occ -> + { { q_onhyps = Some hl; q_concl_occs = occs } } + | hl = LIST0 hypident_occ SEP "," -> + { { q_onhyps = Some hl; q_concl_occs = CAst.make ~loc QNoOccurrences } } + ] ] + ; + clause: + [ [ "in"; cl = in_clause -> { CAst.make ~loc @@ cl } + | "at"; occs = occs_nums -> + { CAst.make ~loc @@ { q_onhyps = Some []; q_concl_occs = occs } } + ] ] + ; + q_clause: + [ [ cl = clause -> { cl } ] ] + ; + concl_occ: + [ [ "*"; occs = occs -> { occs } + | -> { CAst.make ~loc QNoOccurrences } + ] ] + ; + induction_clause: + [ [ c = destruction_arg; pat = as_or_and_ipat; eq = eqn_ipat; + cl = OPT clause -> + { CAst.make ~loc @@ { + indcl_arg = c; + indcl_eqn = eq; + indcl_as = pat; + indcl_in = cl; + } } + ] ] + ; + q_induction_clause: + [ [ cl = induction_clause -> { cl } ] ] + ; + conversion: + [ [ c = Constr.constr -> + { CAst.make ~loc @@ QConvert c } + | c1 = Constr.constr; "with"; c2 = Constr.constr -> + { CAst.make ~loc @@ QConvertWith (c1, c2) } + ] ] + ; + q_conversion: + [ [ c = conversion -> { c } ] ] + ; + orient: + [ [ "->" -> { CAst.make ~loc (Some true) } + | "<-" -> { CAst.make ~loc (Some false) } + | -> { CAst.make ~loc None } + ]] + ; + rewriter: + [ [ "!"; c = constr_with_bindings -> + { (CAst.make ~loc @@ QRepeatPlus,c) } + | [ "?" -> { () } | LEFTQMARK -> { () } ]; c = constr_with_bindings -> + { (CAst.make ~loc @@ QRepeatStar,c) } + | n = lnatural; "!"; c = constr_with_bindings -> + { (CAst.make ~loc @@ QPrecisely n,c) } + | n = lnatural; ["?" -> { () } | LEFTQMARK -> { () } ]; c = constr_with_bindings -> + { (CAst.make ~loc @@ QUpTo n,c) } + | n = lnatural; c = constr_with_bindings -> + { (CAst.make ~loc @@ QPrecisely n,c) } + | c = constr_with_bindings -> + { (CAst.make ~loc @@ QPrecisely (CAst.make 1), c) } + ] ] + ; + oriented_rewriter: + [ [ b = orient; r = rewriter -> + { let (m, c) = r in + CAst.make ~loc @@ { + rew_orient = b; + rew_repeat = m; + rew_equatn = c; + } } + ] ] + ; + q_rewriting: + [ [ r = oriented_rewriter -> { r } ] ] + ; + tactic_then_last: + [ [ "|"; lta = LIST0 (OPT tac2expr LEVEL "6") SEP "|" -> { lta } + | -> { [] } + ] ] + ; + tactic_then_gen: + [ [ ta = tac2expr; "|"; tg = tactic_then_gen -> { let (first,last) = tg in (Some ta :: first, last) } + | ta = tac2expr; ".."; l = tactic_then_last -> { ([], Some (Some ta, l)) } + | ".."; l = tactic_then_last -> { ([], Some (None, l)) } + | ta = tac2expr -> { ([Some ta], None) } + | "|"; tg = tactic_then_gen -> { let (first,last) = tg in (None :: first, last) } + | -> { ([None], None) } + ] ] + ; + q_dispatch: + [ [ d = tactic_then_gen -> { CAst.make ~loc d } ] ] + ; + q_occurrences: + [ [ occs = occs -> { occs } ] ] + ; + red_flag: + [ [ IDENT "beta" -> { CAst.make ~loc @@ QBeta } + | IDENT "iota" -> { CAst.make ~loc @@ QIota } + | IDENT "match" -> { CAst.make ~loc @@ QMatch } + | IDENT "fix" -> { CAst.make ~loc @@ QFix } + | IDENT "cofix" -> { CAst.make ~loc @@ QCofix } + | IDENT "zeta" -> { CAst.make ~loc @@ QZeta } + | IDENT "delta"; d = delta_flag -> { d } + ] ] + ; + refglobal: + [ [ "&"; id = Prim.ident -> { QExpr (CAst.make ~loc @@ QHypothesis id) } + | qid = Prim.qualid -> { QExpr (CAst.make ~loc @@ QReference qid) } + | "$"; id = Prim.ident -> { QAnti (CAst.make ~loc id) } + ] ] + ; + q_reference: + [ [ r = refglobal -> { r } ] ] + ; + refglobals: + [ [ gl = LIST1 refglobal -> { CAst.make ~loc gl } ] ] + ; + delta_flag: + [ [ "-"; "["; idl = refglobals; "]" -> { CAst.make ~loc @@ QDeltaBut idl } + | "["; idl = refglobals; "]" -> { CAst.make ~loc @@ QConst idl } + | -> { CAst.make ~loc @@ QDeltaBut (CAst.make ~loc []) } + ] ] + ; + strategy_flag: + [ [ s = LIST1 red_flag -> { CAst.make ~loc s } + | d = delta_flag -> + { CAst.make ~loc + [CAst.make ~loc QBeta; CAst.make ~loc QIota; CAst.make ~loc QZeta; d] } + ] ] + ; + q_strategy_flag: + [ [ flag = strategy_flag -> { flag } ] ] + ; + hintdb: + [ [ "*" -> { CAst.make ~loc @@ QHintAll } + | l = LIST1 ident_or_anti -> { CAst.make ~loc @@ QHintDbs l } + ] ] + ; + q_hintdb: + [ [ db = hintdb -> { db } ] ] + ; + match_pattern: + [ [ IDENT "context"; id = OPT Prim.ident; + "["; pat = Constr.lconstr_pattern; "]" -> { CAst.make ~loc @@ QConstrMatchContext (id, pat) } + | pat = Constr.lconstr_pattern -> { CAst.make ~loc @@ QConstrMatchPattern pat } ] ] + ; + match_rule: + [ [ mp = match_pattern; "=>"; tac = tac2expr -> + { CAst.make ~loc @@ (mp, tac) } + ] ] + ; + match_list: + [ [ mrl = LIST1 match_rule SEP "|" -> { CAst.make ~loc @@ mrl } + | "|"; mrl = LIST1 match_rule SEP "|" -> { CAst.make ~loc @@ mrl } ] ] + ; + q_constr_matching: + [ [ m = match_list -> { m } ] ] + ; + gmatch_hyp_pattern: + [ [ na = Prim.name; ":"; pat = match_pattern -> { (na, pat) } ] ] + ; + gmatch_pattern: + [ [ "["; hl = LIST0 gmatch_hyp_pattern SEP ","; "|-"; p = match_pattern; "]" -> + { CAst.make ~loc @@ { + q_goal_match_concl = p; + q_goal_match_hyps = hl; + } } + ] ] + ; + gmatch_rule: + [ [ mp = gmatch_pattern; "=>"; tac = tac2expr -> + { CAst.make ~loc @@ (mp, tac) } + ] ] + ; + gmatch_list: + [ [ mrl = LIST1 gmatch_rule SEP "|" -> { CAst.make ~loc @@ mrl } + | "|"; mrl = LIST1 gmatch_rule SEP "|" -> { CAst.make ~loc @@ mrl } ] ] + ; + q_goal_matching: + [ [ m = gmatch_list -> { m } ] ] + ; + move_location: + [ [ "at"; IDENT "top" -> { CAst.make ~loc @@ QMoveFirst } + | "at"; IDENT "bottom" -> { CAst.make ~loc @@ QMoveLast } + | IDENT "after"; id = ident_or_anti -> { CAst.make ~loc @@ QMoveAfter id } + | IDENT "before"; id = ident_or_anti -> { CAst.make ~loc @@ QMoveBefore id } + ] ] + ; + q_move_location: + [ [ mv = move_location -> { mv } ] ] + ; + as_name: + [ [ -> { None } + | "as"; id = ident_or_anti -> { Some id } + ] ] + ; + pose: + [ [ test_lpar_id_coloneq; "("; id = ident_or_anti; ":="; c = Constr.lconstr; ")" -> + { CAst.make ~loc (Some id, c) } + | c = Constr.constr; na = as_name -> { CAst.make ~loc (na, c) } + ] ] + ; + q_pose: + [ [ p = pose -> { p } ] ] + ; + as_ipat: + [ [ "as"; ipat = simple_intropattern -> { Some ipat } + | -> { None } + ] ] + ; + by_tactic: + [ [ "by"; tac = tac2expr -> { Some tac } + | -> { None } + ] ] + ; + assertion: + [ [ test_lpar_id_coloneq; "("; id = ident_or_anti; ":="; c = Constr.lconstr; ")" -> + { CAst.make ~loc (QAssertValue (id, c)) } + | test_lpar_id_colon; "("; id = ident_or_anti; ":"; c = Constr.lconstr; ")"; tac = by_tactic -> + { let ipat = CAst.make ~loc @@ QIntroNaming (CAst.make ~loc @@ QIntroIdentifier id) in + CAst.make ~loc (QAssertType (Some ipat, c, tac)) } + | c = Constr.constr; ipat = as_ipat; tac = by_tactic -> + { CAst.make ~loc (QAssertType (ipat, c, tac)) } + ] ] + ; + q_assert: + [ [ a = assertion -> { a } ] ] + ; +END + +(** Extension of constr syntax *) + +(* +GRAMMAR EXTEND Gram + Pcoq.Constr.operconstr: LEVEL "0" + [ [ IDENT "ltac2"; ":"; "("; tac = tac2expr; ")" -> + { let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in + CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) } + | test_ampersand_ident; "&"; id = Prim.ident -> + { let tac = Tac2quote.of_exact_hyp ~loc (CAst.make ~loc id) in + let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in + CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) } + | test_dollar_ident; "$"; id = Prim.ident -> + { let id = Loc.tag ~loc id in + let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2_quotation) id in + CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) } + ] ] + ; +END +*) +{ + +let () = + +let open Extend in +let open Tok in +let (++) r s = Next (r, s) in +let rules = [ + Rule ( + Stop ++ Aentry test_dollar_ident ++ Atoken (KEYWORD "$") ++ Aentry Prim.ident, + begin fun id _ _ loc -> + let id = Loc.tag ~loc id in + let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2_quotation) id in + CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) + end + ); + + Rule ( + Stop ++ Aentry test_ampersand_ident ++ Atoken (KEYWORD "&") ++ Aentry Prim.ident, + begin fun id _ _ loc -> + let tac = Tac2quote.of_exact_hyp ~loc (CAst.make ~loc id) in + let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in + CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) + end + ); + + Rule ( + Stop ++ Atoken (IDENT "ltac2") ++ Atoken (KEYWORD ":") ++ + Atoken (KEYWORD "(") ++ Aentry tac2expr ++ Atoken (KEYWORD ")"), + begin fun _ tac _ _ _ loc -> + let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in + CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) + end + ) +] in + +Hook.set Tac2entries.register_constr_quotations begin fun () -> + Gram.gram_extend Pcoq.Constr.operconstr (Some (Level "0"), [(None, None, rules)]) +end + +} + +{ + +let pr_ltac2entry _ = mt () (** FIXME *) +let pr_ltac2expr _ = mt () (** FIXME *) + +} + +VERNAC ARGUMENT EXTEND ltac2_entry +PRINTED BY { pr_ltac2entry } +| [ tac2def_val(v) ] -> { v } +| [ tac2def_typ(t) ] -> { t } +| [ tac2def_ext(e) ] -> { e } +| [ tac2def_syn(e) ] -> { e } +| [ tac2def_mut(e) ] -> { e } +| [ tac2def_run(e) ] -> { e } +END + +{ + +let classify_ltac2 = function +| StrSyn _ -> Vernacexpr.VtUnknown, Vernacexpr.VtNow +| StrMut _ | StrVal _ | StrPrm _ | StrTyp _ | StrRun _ -> Vernac_classifier.classify_as_sideeff + +} + +VERNAC COMMAND EXTEND VernacDeclareTactic2Definition +| #[ local = locality ] [ "Ltac2" ltac2_entry(e) ] => { classify_ltac2 e } -> { + Tac2entries.register_struct ?local e + } +END + +{ + +let _ = + let mode = { + Proof_global.name = "Ltac2"; + set = (fun () -> Pvernac.set_command_entry tac2mode); + reset = (fun () -> Pvernac.(set_command_entry Vernac_.noedit_mode)); + } in + Proof_global.register_proof_mode mode + +} + +VERNAC ARGUMENT EXTEND ltac2_expr +PRINTED BY { pr_ltac2expr } +| [ tac2expr(e) ] -> { e } +END + +{ + +open G_ltac +open Vernac_classifier + +} + +VERNAC { tac2mode } EXTEND VernacLtac2 +| [ ltac2_expr(t) ltac_use_default(default) ] => + { classify_as_proofstep } -> { +(* let g = Option.default (Proof_global.get_default_goal_selector ()) g in *) + Tac2entries.call ~default t + } +END + +{ + +open Stdarg + +} + +VERNAC COMMAND EXTEND Ltac2Print CLASSIFIED AS SIDEFF +| [ "Print" "Ltac2" reference(tac) ] -> { Tac2entries.print_ltac tac } +END -- cgit v1.2.3 From f26974c7844d6a58d22c3a9d52b93c5a94f19214 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 29 Oct 2018 11:37:10 +0100 Subject: Fix compilation w.r.t. coq/coq#8779. --- src/tac2core.ml | 2 -- src/tac2tactics.ml | 1 - 2 files changed, 3 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 7384652216..ec10c335e9 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -864,7 +864,6 @@ let constr_flags () = { use_typeclasses = true; solve_unification_constraints = true; - use_hook = Pfedit.solve_by_implicit_tactic (); fail_evar = true; expand_evars = true } @@ -874,7 +873,6 @@ let open_constr_no_classes_flags () = { use_typeclasses = false; solve_unification_constraints = true; - use_hook = Pfedit.solve_by_implicit_tactic (); fail_evar = false; expand_evars = true } diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index 25431af2ea..b06427bc38 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -21,7 +21,6 @@ let thaw r f = Tac2ffi.app_fun1 f Tac2ffi.unit r () let tactic_infer_flags with_evar = { Pretyping.use_typeclasses = true; Pretyping.solve_unification_constraints = true; - Pretyping.use_hook = None; Pretyping.fail_evar = not with_evar; Pretyping.expand_evars = true } -- cgit v1.2.3 From dcdd460710c36522ff10f1b12bbb0b0628c5542f Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Sat, 17 Nov 2018 17:11:20 +0100 Subject: [coq] Overlay to adapt to coq/coq#9003 To be merged when the Coq developers merge the PR upstream. --- src/g_ltac2.mlg | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/g_ltac2.mlg b/src/g_ltac2.mlg index 5aad77596d..2e4fad73fd 100644 --- a/src/g_ltac2.mlg +++ b/src/g_ltac2.mlg @@ -882,8 +882,8 @@ END { let classify_ltac2 = function -| StrSyn _ -> Vernacexpr.VtUnknown, Vernacexpr.VtNow -| StrMut _ | StrVal _ | StrPrm _ | StrTyp _ | StrRun _ -> Vernac_classifier.classify_as_sideeff +| StrSyn _ -> Vernacextend.(VtUnknown, VtNow) +| StrMut _ | StrVal _ | StrPrm _ | StrTyp _ | StrRun _ -> Vernacextend.classify_as_sideeff } @@ -913,7 +913,7 @@ END { open G_ltac -open Vernac_classifier +open Vernacextend } -- cgit v1.2.3 From a67daa71f98cc01a61c4dae4e3dd6bcbf42bb9d4 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Sat, 17 Nov 2018 18:00:56 +0100 Subject: [ci] Add travis setup, docker-based. Pretty straightforward; need to be enabled in the Travis console tho. --- .travis.yml | 39 +++++++++++++++++++++++++++++++++++++++ 1 file changed, 39 insertions(+) create mode 100644 .travis.yml diff --git a/.travis.yml b/.travis.yml new file mode 100644 index 0000000000..d52b3ebd25 --- /dev/null +++ b/.travis.yml @@ -0,0 +1,39 @@ +dist: trusty +sudo: required +language: generic + +services: + - docker + +env: + global: + - NJOBS="2" + - CONTRIB_NAME="ltac2" + matrix: + - COQ_IMAGE="coqorg/coq:dev" + +install: | + # Prepare the COQ container + docker run -d -i --init --name=COQ -v ${TRAVIS_BUILD_DIR}:/home/coq/${CONTRIB_NAME} -w /home/coq/${CONTRIB_NAME} ${COQ_IMAGE} + docker exec COQ /bin/bash --login -c " + # This bash script is double-quoted to interpolate Travis CI env vars: + echo \"Build triggered by ${TRAVIS_EVENT_TYPE}\" + export PS4='+ \e[33;1m(\$0 @ line \$LINENO) \$\e[0m ' + set -ex # -e = exit on failure; -x = trace for debug + # opam update -y + # opam install -y -j ${NJOBS} coq-mathcomp-ssreflect + opam config list + opam repo list + opam list + " +script: +- echo -e "${ANSI_YELLOW}Building ${CONTRIB_NAME}...${ANSI_RESET}" && echo -en 'travis_fold:start:script\\r' +- | + docker exec COQ /bin/bash --login -c " + export PS4='+ \e[33;1m(\$0 @ line \$LINENO) \$\e[0m ' + set -ex + sudo chown -R coq:coq /home/coq/${CONTRIB_NAME} + make + " +- docker stop COQ # optional +- echo -en 'travis_fold:end:script\\r' -- cgit v1.2.3 From fd184924e1d8955d6cfe7d7645dfb8776b211195 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 17 Nov 2018 19:20:29 +0100 Subject: Add an image status for the CI. --- README.md | 2 ++ 1 file changed, 2 insertions(+) diff --git a/README.md b/README.md index 73785f6368..d49dd88076 100644 --- a/README.md +++ b/README.md @@ -1,3 +1,5 @@ +[![Build Status](https://travis-ci.org/ppedrot/ltac2.svg?branch=master)](https://travis-ci.org/ppedrot/ltac2) + Overview ======== -- cgit v1.2.3 From 3ba8647971c441307dd61bc67dc2c3705b345b56 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 19 Nov 2018 09:51:05 +0100 Subject: Add a Char module. --- _CoqProject | 1 + src/tac2core.ml | 10 ++++++++++ theories/Char.v | 12 ++++++++++++ theories/Ltac2.v | 1 + 4 files changed, 24 insertions(+) create mode 100644 theories/Char.v diff --git a/_CoqProject b/_CoqProject index 071066dd86..e2ef5cebe1 100644 --- a/_CoqProject +++ b/_CoqProject @@ -35,6 +35,7 @@ src/ltac2_plugin.mlpack theories/Init.v theories/Int.v +theories/Char.v theories/String.v theories/Ident.v theories/Array.v diff --git a/src/tac2core.ml b/src/tac2core.ml index ec10c335e9..b6983ed869 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -258,6 +258,16 @@ let () = define1 "int_neg" int begin fun m -> return (Value.of_int (~- m)) end +(** Char *) + +let () = define1 "char_of_int" int begin fun n -> + wrap (fun () -> Value.of_char (Char.chr n)) +end + +let () = define1 "char_to_int" char begin fun n -> + wrap (fun () -> Value.of_int (Char.code n)) +end + (** String *) let () = define2 "string_make" int char begin fun n c -> diff --git a/theories/Char.v b/theories/Char.v new file mode 100644 index 0000000000..29fef60f2c --- /dev/null +++ b/theories/Char.v @@ -0,0 +1,12 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* char := "ltac2" "char_of_int". +Ltac2 @external to_int : char -> int := "ltac2" "char_to_int". diff --git a/theories/Ltac2.v b/theories/Ltac2.v index 3fe71f4c65..e838fb7b81 100644 --- a/theories/Ltac2.v +++ b/theories/Ltac2.v @@ -9,6 +9,7 @@ Require Export Ltac2.Init. Require Ltac2.Int. +Require Ltac2.Char. Require Ltac2.String. Require Ltac2.Ident. Require Ltac2.Array. -- cgit v1.2.3 From 93300e662b6e7571619508e6f6d47b963d5300d1 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 19 Nov 2018 10:00:50 +0100 Subject: Add a function to generate fresh reference instances. --- src/tac2core.ml | 8 ++++++++ theories/Env.v | 5 +++++ 2 files changed, 13 insertions(+) diff --git a/src/tac2core.ml b/src/tac2core.ml index b6983ed869..890062a6d1 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -867,6 +867,14 @@ let () = define1 "env_path" reference begin fun r -> throw err_notfound end +let () = define1 "env_instantiate" reference begin fun r -> + Proofview.tclENV >>= fun env -> + Proofview.tclEVARMAP >>= fun sigma -> + let (sigma, c) = Evd.fresh_global env sigma r in + Proofview.Unsafe.tclEVARS sigma >>= fun () -> + return (Value.of_constr c) +end + (** ML types *) let constr_flags () = diff --git a/theories/Env.v b/theories/Env.v index 7e36aa7990..c9b250f4ba 100644 --- a/theories/Env.v +++ b/theories/Env.v @@ -20,3 +20,8 @@ Ltac2 @ external expand : ident list -> Std.reference list := "ltac2" "env_expan Ltac2 @ external path : Std.reference -> ident list := "ltac2" "env_path". (** Returns the absolute name of the given reference. Panics if the reference does not exist. *) + +Ltac2 @ external instantiate : Std.reference -> constr := "ltac2" "env_instantiate". +(** Returns a fresh instance of the corresponding reference, in particular + generating fresh universe variables and constraints when this reference is + universe-polymorphic. *) -- cgit v1.2.3 From 387a56ced3a093af1e97ed08be02c93ceaf66aa8 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 19 Nov 2018 10:11:44 +0100 Subject: Adding a module to manipulate Ltac1 values. --- _CoqProject | 1 + doc/ltac2.md | 17 +++++++++ src/g_ltac2.mlg | 2 ++ src/tac2core.ml | 103 ++++++++++++++++++++++++++++++++++++++++++++++++++++-- src/tac2env.ml | 3 ++ src/tac2env.mli | 3 ++ src/tac2ffi.ml | 1 + src/tac2ffi.mli | 1 + src/tac2quote.ml | 1 + src/tac2quote.mli | 4 +++ theories/Ltac1.v | 36 +++++++++++++++++++ theories/Ltac2.v | 1 + 12 files changed, 171 insertions(+), 2 deletions(-) create mode 100644 theories/Ltac1.v diff --git a/_CoqProject b/_CoqProject index e2ef5cebe1..dda5a8001a 100644 --- a/_CoqProject +++ b/_CoqProject @@ -47,4 +47,5 @@ theories/Fresh.v theories/Std.v theories/Env.v theories/Notations.v +theories/Ltac1.v theories/Ltac2.v diff --git a/doc/ltac2.md b/doc/ltac2.md index 3cee0ac494..b217cb08e6 100644 --- a/doc/ltac2.md +++ b/doc/ltac2.md @@ -880,6 +880,8 @@ a backtrace. ## Ltac1 from Ltac2 +### Simple API + One can call Ltac1 code from Ltac2 by using the `ltac1` quotation. It parses a Ltac1 expression, and semantics of this quotation is the evaluation of the corresponding code for its side effects. In particular, in cannot return values, @@ -888,6 +890,21 @@ and the quotation has type `unit`. Beware, Ltac1 **cannot** access variables from the Ltac2 scope. One is limited to the use of standalone function calls. +### Low-level API + +There exists a lower-level FFI into Ltac1 that is not recommended for daily use, +which is available in the `Ltac2.Ltac1` module. This API allows to directly +manipulate dynamically-typed Ltac1 values, either through the function calls, +or using the `ltac1val` quotation. The latter parses the same as `ltac1`, but +has type `Ltac2.Ltac1.t` instead of `unit`, and dynamically behaves as an Ltac1 +thunk, i.e. `ltac1val:(foo)` corresponds to the tactic closure that Ltac1 +would generate from `idtac; foo`. + +Due to intricate dynamic semantics, understanding when Ltac1 value quotations +focus is very hard. This is why some functions return a continuation-passing +style value, as it can dispatch dynamically between focused and unfocused +behaviour. + ## Ltac2 from Ltac1 Same as above by switching Ltac1 by Ltac2 and using the `ltac2` quotation diff --git a/src/g_ltac2.mlg b/src/g_ltac2.mlg index 5aad77596d..0494227f1e 100644 --- a/src/g_ltac2.mlg +++ b/src/g_ltac2.mlg @@ -100,6 +100,7 @@ let inj_open_constr loc c = inj_wit Tac2quote.wit_open_constr loc c let inj_pattern loc c = inj_wit Tac2quote.wit_pattern loc c let inj_reference loc c = inj_wit Tac2quote.wit_reference loc c let inj_ltac1 loc e = inj_wit Tac2quote.wit_ltac1 loc e +let inj_ltac1val loc e = inj_wit Tac2quote.wit_ltac1val loc e let pattern_of_qualid qid = if Tac2env.is_constructor qid then CAst.make ?loc:qid.CAst.loc @@ CPatRef (RelId qid, []) @@ -224,6 +225,7 @@ GRAMMAR EXTEND Gram | IDENT "pattern"; ":"; "("; c = Constr.lconstr_pattern; ")" -> { inj_pattern loc c } | IDENT "reference"; ":"; "("; c = globref; ")" -> { inj_reference loc c } | IDENT "ltac1"; ":"; "("; qid = ltac1_expr; ")" -> { inj_ltac1 loc qid } + | IDENT "ltac1val"; ":"; "("; qid = ltac1_expr; ")" -> { inj_ltac1val loc qid } ] ] ; let_clause: diff --git a/src/tac2core.ml b/src/tac2core.ml index 890062a6d1..aad4814744 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -20,8 +20,11 @@ open Proofview.Notations module Value = Tac2ffi open Value -let std_core n = KerName.make Tac2env.std_prefix (Label.of_id (Id.of_string_soft n)) -let coq_core n = KerName.make Tac2env.coq_prefix (Label.of_id (Id.of_string_soft n)) +let core_prefix path n = KerName.make path (Label.of_id (Id.of_string_soft n)) + +let std_core n = core_prefix Tac2env.std_prefix n +let coq_core n = core_prefix Tac2env.coq_prefix n +let ltac1_core n = core_prefix Tac2env.ltac1_prefix n module Core = struct @@ -37,6 +40,7 @@ let t_ident = coq_core "ident" let t_option = coq_core "option" let t_exn = coq_core "exn" let t_reference = std_core "reference" +let t_ltac1 = ltac1_core "t" let c_nil = coq_core "[]" let c_cons = coq_core "::" @@ -875,6 +879,73 @@ let () = define1 "env_instantiate" reference begin fun r -> return (Value.of_constr c) end +(** Ltac1 in Ltac2 *) + +let ltac1 = Tac2ffi.repr_ext Value.val_ltac1 +let of_ltac1 v = Value.of_ext Value.val_ltac1 v + +let () = define1 "ltac1_ref" (list ident) begin fun ids -> + let open Ltac_plugin in + let r = match ids with + | [] -> raise Not_found + | _ :: _ as ids -> + let (id, path) = List.sep_last ids in + let path = DirPath.make (List.rev path) in + let fp = Libnames.make_path path id in + if Tacenv.exists_tactic fp then + List.hd (Tacenv.locate_extended_all_tactic (Libnames.qualid_of_path fp)) + else raise Not_found + in + let tac = Tacinterp.Value.of_closure (Tacinterp.default_ist ()) (Tacenv.interp_ltac r) in + return (Value.of_ext val_ltac1 tac) +end + +let () = define1 "ltac1_run" ltac1 begin fun v -> + let open Ltac_plugin in + Tacinterp.tactic_of_value (Tacinterp.default_ist ()) v >>= fun () -> + return v_unit +end + +let () = define3 "ltac1_apply" ltac1 (list ltac1) closure begin fun f args k -> + let open Ltac_plugin in + let open Tacexpr in + let open Locus in + let k ret = + Proofview.tclIGNORE (Tac2ffi.apply k [Value.of_ext val_ltac1 ret]) + in + let fold arg (i, vars, lfun) = + let id = Id.of_string ("x" ^ string_of_int i) in + let x = Reference (ArgVar CAst.(make 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 + let tac = TacArg(CAst.make @@ TacCall (CAst.make (ArgVar CAst.(make @@ Id.of_string "F"),args))) in + Tacinterp.val_interp ist tac k >>= fun () -> + return v_unit +end + +let () = define1 "ltac1_of_constr" constr begin fun c -> + let open Ltac_plugin in + return (Value.of_ext val_ltac1 (Tacinterp.Value.of_constr c)) +end + +let () = define1 "ltac1_to_constr" ltac1 begin fun v -> + let open Ltac_plugin in + return (Value.of_option Value.of_constr (Tacinterp.Value.to_constr v)) +end + +let () = define1 "ltac1_of_list" (list ltac1) begin fun l -> + let open Geninterp.Val in + return (Value.of_ext val_ltac1 (inject (Base typ_list) l)) +end + +let () = define1 "ltac1_to_list" ltac1 begin fun v -> + let open Ltac_plugin in + return (Value.of_option (Value.of_list of_ltac1) (Tacinterp.Value.to_list v)) +end + (** ML types *) let constr_flags () = @@ -1037,6 +1108,34 @@ let () = } in define_ml_object Tac2quote.wit_ltac1 obj +let () = + let open Ltac_plugin in + let intern self ist tac = + (** Prevent inner calls to Ltac2 values *) + let extra = Tac2intern.drop_ltac2_env ist.Genintern.extra in + let ist = { ist with Genintern.extra } in + let _, tac = Genintern.intern Ltac_plugin.Tacarg.wit_tactic ist tac in + GlbVal tac, gtypref t_ltac1 + in + let interp ist tac = + let ist = { env_ist = Id.Map.empty } in + let lfun = Tac2interp.set_env ist Id.Map.empty in + let ist = Ltac_plugin.Tacinterp.default_ist () in + let ist = { ist with Geninterp.lfun = lfun } in + return (Value.of_ext val_ltac1 (Tacinterp.Value.of_closure ist tac)) + in + let subst s tac = Genintern.substitute Tacarg.wit_tactic s tac in + let print env tac = + str "ltac1val:(" ++ Ltac_plugin.Pptactic.pr_glob_tactic env tac ++ str ")" + in + let obj = { + ml_intern = intern; + ml_subst = subst; + ml_interp = interp; + ml_print = print; + } in + define_ml_object Tac2quote.wit_ltac1val obj + (** Ltac2 in terms *) let () = diff --git a/src/tac2env.ml b/src/tac2env.ml index dcf7440498..8198f92ff8 100644 --- a/src/tac2env.ml +++ b/src/tac2env.ml @@ -276,6 +276,9 @@ let coq_prefix = let std_prefix = MPfile (DirPath.make (List.map Id.of_string ["Std"; "Ltac2"])) +let ltac1_prefix = + MPfile (DirPath.make (List.map Id.of_string ["Ltac1"; "Ltac2"])) + (** Generic arguments *) let wit_ltac2 = Genarg.make0 "ltac2:value" diff --git a/src/tac2env.mli b/src/tac2env.mli index 7616579d63..c7e87c5432 100644 --- a/src/tac2env.mli +++ b/src/tac2env.mli @@ -133,6 +133,9 @@ val coq_prefix : ModPath.t val std_prefix : ModPath.t (** Path where Ltac-specific datatypes are defined in Ltac2 plugin. *) +val ltac1_prefix : ModPath.t +(** Path where the Ltac1 legacy FFI is defined. *) + (** {5 Generic arguments} *) val wit_ltac2 : (raw_tacexpr, glb_tacexpr, Util.Empty.t) genarg_type diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index df1857c3e7..c271967bd6 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -96,6 +96,7 @@ let val_projection = Val.create "projection" let val_case = Val.create "case" let val_univ = Val.create "universe" let val_free : Names.Id.Set.t Val.tag = Val.create "free" +let val_ltac1 : Geninterp.Val.t Val.tag = Val.create "ltac1" let extract_val (type a) (type b) (tag : a Val.tag) (tag' : b Val.tag) (v : b) : a = match Val.eq tag tag' with diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli index d801c4f605..bfc93d99e6 100644 --- a/src/tac2ffi.mli +++ b/src/tac2ffi.mli @@ -167,6 +167,7 @@ val val_projection : Projection.t Val.tag val val_case : Constr.case_info Val.tag val val_univ : Univ.Level.t Val.tag val val_free : Id.Set.t Val.tag +val val_ltac1 : Geninterp.Val.t Val.tag val val_exn : Exninfo.iexn Tac2dyn.Val.tag (** Toplevel representation of OCaml exceptions. Invariant: no [LtacError] diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 3bddfe7594..5a26e7465c 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -22,6 +22,7 @@ let wit_ident = Arg.create "ident" let wit_constr = Arg.create "constr" let wit_open_constr = Arg.create "open_constr" let wit_ltac1 = Arg.create "ltac1" +let wit_ltac1val = Arg.create "ltac1val" (** Syntactic quoting of expressions. *) diff --git a/src/tac2quote.mli b/src/tac2quote.mli index 09aa92f9ee..1b03dad8ec 100644 --- a/src/tac2quote.mli +++ b/src/tac2quote.mli @@ -96,3 +96,7 @@ val wit_constr : (Constrexpr.constr_expr, Glob_term.glob_constr) Arg.tag val wit_open_constr : (Constrexpr.constr_expr, Glob_term.glob_constr) Arg.tag val wit_ltac1 : (Ltac_plugin.Tacexpr.raw_tactic_expr, Ltac_plugin.Tacexpr.glob_tactic_expr) Arg.tag +(** Ltac1 AST quotation, seen as a 'tactic'. Its type is unit in Ltac2. *) + +val wit_ltac1val : (Ltac_plugin.Tacexpr.raw_tactic_expr, Ltac_plugin.Tacexpr.glob_tactic_expr) Arg.tag +(** Ltac1 AST quotation, seen as a value-returning expression, with type Ltac1.t. *) diff --git a/theories/Ltac1.v b/theories/Ltac1.v new file mode 100644 index 0000000000..c4e0b606d0 --- /dev/null +++ b/theories/Ltac1.v @@ -0,0 +1,36 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* t := "ltac2" "ltac1_ref". +(** Returns the Ltac1 definition with the given absolute name. *) + +Ltac2 @ external run : t -> unit := "ltac2" "ltac1_run". +(** Runs an Ltac1 value, assuming it is a 'tactic', i.e. not returning + anything. *) + +Ltac2 @ external apply : t -> t list -> (t -> unit) -> unit := "ltac2" "ltac1_apply". +(** Applies an Ltac1 value to a list of arguments, and provides the result in + CPS style. It does **not** run the returned value. *) + +(** Conversion functions *) + +Ltac2 @ external of_constr : constr -> t := "ltac2" "ltac1_of_constr". +Ltac2 @ external to_constr : t -> constr option := "ltac2" "ltac1_to_constr". + +Ltac2 @ external of_list : t list -> t := "ltac2" "ltac1_of_list". +Ltac2 @ external to_list : t -> t list option := "ltac2" "ltac1_to_list". diff --git a/theories/Ltac2.v b/theories/Ltac2.v index e838fb7b81..ac90f63560 100644 --- a/theories/Ltac2.v +++ b/theories/Ltac2.v @@ -20,4 +20,5 @@ Require Ltac2.Fresh. Require Ltac2.Pattern. Require Ltac2.Std. Require Ltac2.Env. +Require Ltac2.Ltac1. Require Export Ltac2.Notations. -- cgit v1.2.3 From f05f4ebe9e91829e2817c63d90ec328f430992c4 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 23 Nov 2018 14:04:25 +0100 Subject: Fix w.r.t. coq/coq#9051. --- src/g_ltac2.mlg | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/g_ltac2.mlg b/src/g_ltac2.mlg index 2a20264edc..12364a0bd8 100644 --- a/src/g_ltac2.mlg +++ b/src/g_ltac2.mlg @@ -25,7 +25,7 @@ type lookahead = int -> Tok.t Stream.t -> int option let entry_of_lookahead s (lk : lookahead) = let run strm = match lk 0 strm with None -> err () | Some _ -> () in - Gram.Entry.of_parser s run + Pcoq.Entry.of_parser s run let (>>) (lk1 : lookahead) lk2 n strm = match lk1 n strm with | None -> None @@ -859,7 +859,7 @@ let rules = [ ] in Hook.set Tac2entries.register_constr_quotations begin fun () -> - Gram.gram_extend Pcoq.Constr.operconstr (Some (Level "0"), [(None, None, rules)]) + Pcoq.grammar_extend Pcoq.Constr.operconstr None (Some (Level "0"), [(None, None, rules)]) end } -- cgit v1.2.3 From 09b917593e90233d7a55610e9ce7886de77ef576 Mon Sep 17 00:00:00 2001 From: James R. Wilcox Date: Sat, 24 Nov 2018 13:22:50 -0800 Subject: tests/Makefile: support unset COQBIN, like top-level Makefile does --- tests/Makefile | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/tests/Makefile b/tests/Makefile index 9370b063f8..37915e0d6f 100644 --- a/tests/Makefile +++ b/tests/Makefile @@ -1,3 +1,7 @@ +ifeq "$(COQBIN)" "" + COQBIN=$(dir $(shell which coqtop))/ +endif + all: $(patsubst %.v,%.v.log,$(wildcard *.v)) %.v.log: %.v -- cgit v1.2.3 From e2fe373a1947206746dd43afe6c9815c69453def Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Wed, 28 Nov 2018 00:47:10 +0100 Subject: [build] Test tests in Travis, use coqc for tests. `coqtop -batch` is an oxymoron, in prevision for upstream changes use `coqc`. We also call `make test` in Travis as to make CI more robust. --- .travis.yml | 1 + tests/Makefile | 4 ++-- tests/example1.v | 1 + 3 files changed, 4 insertions(+), 2 deletions(-) diff --git a/.travis.yml b/.travis.yml index d52b3ebd25..2628abde45 100644 --- a/.travis.yml +++ b/.travis.yml @@ -34,6 +34,7 @@ script: set -ex sudo chown -R coq:coq /home/coq/${CONTRIB_NAME} make + make tests " - docker stop COQ # optional - echo -en 'travis_fold:end:script\\r' diff --git a/tests/Makefile b/tests/Makefile index 37915e0d6f..d85ae90dd6 100644 --- a/tests/Makefile +++ b/tests/Makefile @@ -1,11 +1,11 @@ ifeq "$(COQBIN)" "" - COQBIN=$(dir $(shell which coqtop))/ + COQBIN=$(dir $(shell which coqc))/ endif all: $(patsubst %.v,%.v.log,$(wildcard *.v)) %.v.log: %.v - $(COQBIN)/coqtop -batch -I ../src -Q ../theories Ltac2 -lv $< > $@ + $(COQBIN)/coqc -I ../src -Q ../theories Ltac2 $< > $@ if [ $$? = 0 ]; then \ echo " $<... OK"; \ else \ diff --git a/tests/example1.v b/tests/example1.v index 1b26aad824..023791050f 100644 --- a/tests/example1.v +++ b/tests/example1.v @@ -24,3 +24,4 @@ Goal forall n m, n + m = 0 -> n = 0. Proof. refine (fun () => '(fun n m H => _)). let t := get_hyp_by_name @H in Message.print (Message.of_constr t). +Abort. -- cgit v1.2.3 From 8b1c09e551c02f26c524922570341f0f7fc78e2e Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Mon, 10 Dec 2018 21:07:45 +0100 Subject: [coq] Adapt to coq/coq#9172 Note that this highlights some issues with the current Coq interface, not clear what should we do. --- src/g_ltac2.mlg | 2 +- src/tac2core.ml | 3 ++- src/tac2entries.ml | 3 ++- 3 files changed, 5 insertions(+), 3 deletions(-) diff --git a/src/g_ltac2.mlg b/src/g_ltac2.mlg index 12364a0bd8..7020ca079e 100644 --- a/src/g_ltac2.mlg +++ b/src/g_ltac2.mlg @@ -859,7 +859,7 @@ let rules = [ ] in Hook.set Tac2entries.register_constr_quotations begin fun () -> - Pcoq.grammar_extend Pcoq.Constr.operconstr None (Some (Level "0"), [(None, None, rules)]) + Pcoq.grammar_extend Pcoq.Constr.operconstr None (Some (Gramlib.Gramext.Level "0"), [(None, None, rules)]) end } diff --git a/src/tac2core.ml b/src/tac2core.ml index aad4814744..8e92e154ac 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -1142,7 +1142,8 @@ let () = let interp ist env sigma concl tac = let ist = Tac2interp.get_env ist in let tac = Proofview.tclIGNORE (Tac2interp.interp ist tac) in - let c, sigma = Pfedit.refine_by_tactic env sigma concl tac in + let name, poly = Id.of_string "ltac2", false in + let c, sigma = Pfedit.refine_by_tactic ~name ~poly env sigma concl tac in (EConstr.of_constr c, sigma) in GlobEnv.register_constr_interp0 wit_ltac2 interp diff --git a/src/tac2entries.ml b/src/tac2entries.ml index bba4680a72..22025f0a8f 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -750,7 +750,8 @@ let perform_eval e = Proof_global.give_me_the_proof () with Proof_global.NoCurrentProof -> let sigma = Evd.from_env env in - Goal_select.SelectAll, Proof.start sigma [] + let name, poly = Id.of_string "ltac2", false in + Goal_select.SelectAll, Proof.start ~name ~poly sigma [] in let v = match selector with | Goal_select.SelectNth i -> Proofview.tclFOCUS i i v -- cgit v1.2.3 From c9faee36005bea6add36b0eadb87af0f7439bb41 Mon Sep 17 00:00:00 2001 From: Maxime Dénès Date: Thu, 17 Jan 2019 12:27:36 +0100 Subject: Adapt to Coq's new proof mode API --- src/g_ltac2.mlg | 8 +------- 1 file changed, 1 insertion(+), 7 deletions(-) diff --git a/src/g_ltac2.mlg b/src/g_ltac2.mlg index 7020ca079e..c707a82e5b 100644 --- a/src/g_ltac2.mlg +++ b/src/g_ltac2.mlg @@ -897,13 +897,7 @@ END { -let _ = - let mode = { - Proof_global.name = "Ltac2"; - set = (fun () -> Pvernac.set_command_entry tac2mode); - reset = (fun () -> Pvernac.(set_command_entry Vernac_.noedit_mode)); - } in - Proof_global.register_proof_mode mode +let _ = Pvernac.register_proof_mode "Ltac2" tac2mode } -- cgit v1.2.3 From 49155a0817234299c45d04a14bd834f44fbc391f Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Mon, 28 Jan 2019 18:00:02 -0500 Subject: Make lazy_match! goal actually lazy It was missing `Control.once`. Fixes coq/ltac2#79 Fixes coq/ltac2#77 --- tests/matching.v | 19 +++++++++++++++++++ theories/Pattern.v | 2 +- 2 files changed, 20 insertions(+), 1 deletion(-) diff --git a/tests/matching.v b/tests/matching.v index 6bc5706da7..4338cbd32f 100644 --- a/tests/matching.v +++ b/tests/matching.v @@ -50,3 +50,22 @@ match! reverse goal with check_id h' @i end. Abort. + +(* Check #79 *) +Goal 2 = 3. + Control.plus + (fun () + => lazy_match! goal with + | [ |- 2 = 3 ] => Control.zero (Tactic_failure None) + | [ |- 2 = _ ] => Control.zero (Tactic_failure (Some (Message.of_string "should not be printed"))) + end) + (fun e + => match e with + | Tactic_failure c + => match c with + | None => () + | _ => Control.zero e + end + | e => Control.zero e + end). +Abort. diff --git a/theories/Pattern.v b/theories/Pattern.v index ff7776b682..8d1fb0cd8a 100644 --- a/theories/Pattern.v +++ b/theories/Pattern.v @@ -125,7 +125,7 @@ Ltac2 lazy_goal_match0 rev pats := in Control.plus cur next end in - interp pats (). + Control.once (fun () => interp pats) (). Ltac2 multi_goal_match0 rev pats := let rec interp m := match m with -- cgit v1.2.3 From 14cb3c26e5b35a4d824838c76a7cf8d8a0fa35e0 Mon Sep 17 00:00:00 2001 From: Maxime Dénès Date: Fri, 1 Feb 2019 20:59:01 +0100 Subject: Adapt to https://github.com/coq/coq/pull/9410 --- src/tac2core.ml | 6 ++++-- src/tac2tactics.ml | 4 +++- 2 files changed, 7 insertions(+), 3 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index 8e92e154ac..b5ae446ed5 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -954,7 +954,8 @@ let constr_flags () = use_typeclasses = true; solve_unification_constraints = true; fail_evar = true; - expand_evars = true + expand_evars = true; + program_mode = false; } let open_constr_no_classes_flags () = @@ -963,7 +964,8 @@ let open_constr_no_classes_flags () = use_typeclasses = false; solve_unification_constraints = true; fail_evar = false; - expand_evars = true + expand_evars = true; + program_mode = false; } (** Embed all Ltac2 data into Values *) diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index b06427bc38..bc92ab43a8 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -22,7 +22,9 @@ let tactic_infer_flags with_evar = { Pretyping.use_typeclasses = true; Pretyping.solve_unification_constraints = true; Pretyping.fail_evar = not with_evar; - Pretyping.expand_evars = true } + Pretyping.expand_evars = true; + Pretyping.program_mode = false; +} (** FIXME: export a better interface in Tactics *) let delayed_of_tactic tac env sigma = -- cgit v1.2.3 From 8f81a501b2dacd779b0c0eb75091ea10bdb8f2d7 Mon Sep 17 00:00:00 2001 From: Maxime Dénès Date: Fri, 8 Feb 2019 22:14:13 +0100 Subject: Remove VtUnknown classification That classification is going to disappear from Coq. However, I don't understand why it was used here. Can you confirm that the command can not open a proof? --- src/g_ltac2.mlg | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/g_ltac2.mlg b/src/g_ltac2.mlg index c707a82e5b..609e8a6c0e 100644 --- a/src/g_ltac2.mlg +++ b/src/g_ltac2.mlg @@ -884,7 +884,7 @@ END { let classify_ltac2 = function -| StrSyn _ -> Vernacextend.(VtUnknown, VtNow) +| StrSyn _ -> Vernacextend.(VtSideff [], VtNow) | StrMut _ | StrVal _ | StrPrm _ | StrTyp _ | StrRun _ -> Vernacextend.classify_as_sideeff } -- cgit v1.2.3 From 67cff8c545a25e7fa1a29b08d41fc64a7278508b Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Mon, 17 Dec 2018 19:01:04 +0100 Subject: [coq] Adapt to coq/coq#9137 To be merged when the upstream PR is merged. Not sure this is the right thing to do tho. --- src/dune | 1 + src/tac2tactics.ml | 3 ++- 2 files changed, 3 insertions(+), 1 deletion(-) diff --git a/src/dune b/src/dune index 7c911fb041..4a018adb9a 100644 --- a/src/dune +++ b/src/dune @@ -2,6 +2,7 @@ (name ltac2) (public_name coq.plugins.ltac2) (modules_without_implementation tac2expr tac2qexpr tac2types) + (flags :standard -w -50) (libraries coq.plugins.firstorder)) (rule diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index bc92ab43a8..059a1babd7 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -29,7 +29,8 @@ let tactic_infer_flags with_evar = { (** FIXME: export a better interface in Tactics *) let delayed_of_tactic tac env sigma = let _, pv = Proofview.init sigma [] in - let c, pv, _, _ = Proofview.apply env tac pv in + let name, poly = Id.of_string "ltac2_delayed", false in + let c, pv, _, _ = Proofview.apply ~name ~poly env tac pv in (sigma, c) let delayed_of_thunk r tac env sigma = -- cgit v1.2.3 From 95483808fa2f95b3ef8fc6b3b6da14c23c88d620 Mon Sep 17 00:00:00 2001 From: Gaëtan Gilbert Date: Wed, 19 Dec 2018 16:30:12 +0100 Subject: Adapt to coq/coq#8817 (SProp) --- src/tac2core.ml | 49 +++++++++++++++++++++++++++++++++++-------------- 1 file changed, 35 insertions(+), 14 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index b5ae446ed5..78cbe6d2be 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -66,6 +66,26 @@ let to_name c = match Value.to_option Value.to_ident c with | None -> Anonymous | Some id -> Name id +let of_relevance = function + | Sorts.Relevant -> ValInt 0 + | Sorts.Irrelevant -> ValInt 1 + +let to_relevance = function + | ValInt 0 -> Sorts.Relevant + | ValInt 1 -> Sorts.Irrelevant + | _ -> assert false + +let of_annot f Context.{binder_name;binder_relevance} = + of_tuple [|(f binder_name); of_relevance binder_relevance|] + +let to_annot f x = + match to_tuple x with + | [|x;y|] -> + let x = f x in + let y = to_relevance y in + Context.make_annot x y + | _ -> assert false + let of_instance u = let u = Univ.Instance.to_array (EConstr.Unsafe.to_instance u) in Value.of_array (fun v -> Value.of_ext Value.val_univ v) u @@ -75,12 +95,12 @@ let to_instance u = EConstr.EInstance.make (Univ.Instance.of_array u) let of_rec_declaration (nas, ts, cs) = - (Value.of_array of_name nas, + (Value.of_array (of_annot of_name) nas, Value.of_array Value.of_constr ts, Value.of_array Value.of_constr cs) let to_rec_declaration (nas, ts, cs) = - (Value.to_array to_name nas, + (Value.to_array (to_annot to_name) nas, Value.to_array Value.to_constr ts, Value.to_array Value.to_constr cs) @@ -338,19 +358,19 @@ let () = define1 "constr_kind" constr begin fun c -> |] | Prod (na, t, u) -> v_blk 6 [| - of_name na; + of_annot of_name na; Value.of_constr t; Value.of_constr u; |] | Lambda (na, t, c) -> v_blk 7 [| - of_name na; + of_annot of_name na; Value.of_constr t; Value.of_constr c; |] | LetIn (na, b, t, c) -> v_blk 8 [| - of_name na; + of_annot of_name na; Value.of_constr b; Value.of_constr t; Value.of_constr c; @@ -431,17 +451,17 @@ let () = define1 "constr_make" valexpr begin fun knd -> let t = Value.to_constr t in EConstr.mkCast (c, k, t) | (6, [|na; t; u|]) -> - let na = to_name na in + let na = to_annot to_name na in let t = Value.to_constr t in let u = Value.to_constr u in EConstr.mkProd (na, t, u) | (7, [|na; t; c|]) -> - let na = to_name na in + let na = to_annot to_name na in let t = Value.to_constr t in let u = Value.to_constr c in EConstr.mkLambda (na, t, u) | (8, [|na; b; t; c|]) -> - let na = to_name na in + let na = to_annot to_name na in let b = Value.to_constr b in let t = Value.to_constr t in let c = Value.to_constr c in @@ -511,7 +531,7 @@ end let () = define1 "constr_case" (repr_ext val_inductive) begin fun ind -> Proofview.tclENV >>= fun env -> try - let ans = Inductiveops.make_case_info env ind Constr.RegularStyle in + let ans = Inductiveops.make_case_info env ind Sorts.Relevant Constr.RegularStyle in return (Value.of_ext Value.val_case ans) with e when CErrors.noncritical e -> throw err_notfound @@ -544,7 +564,7 @@ let () = define3 "constr_in_context" ident constr closure begin fun id t c -> Tacticals.New.tclZEROMSG (str "Variable already exists") else let open Context.Named.Declaration in - let nenv = EConstr.push_named (LocalAssum (id, t)) env in + let nenv = EConstr.push_named (LocalAssum (Context.make_annot id Sorts.Relevant, t)) env in let (sigma, (evt, _)) = Evarutil.new_type_evar nenv sigma Evd.univ_flexible in let (sigma, evk) = Evarutil.new_pure_evar (Environ.named_context_val nenv) sigma evt in Proofview.Unsafe.tclEVARS sigma >>= fun () -> @@ -554,7 +574,7 @@ let () = define3 "constr_in_context" ident constr closure begin fun id t c -> let args = List.map (fun d -> EConstr.mkVar (get_id d)) (EConstr.named_context env) in let args = Array.of_list (EConstr.mkRel 1 :: args) in let ans = EConstr.mkEvar (evk, args) in - let ans = EConstr.mkLambda (Name id, t, ans) in + let ans = EConstr.mkLambda (Context.make_annot (Name id) Sorts.Relevant, t, ans) in return (Value.of_constr ans) | _ -> throw err_notfocussed @@ -759,16 +779,17 @@ end let () = define0 "hyps" begin pf_apply begin fun env _ -> - let open Context.Named.Declaration in + let open Context in + let open Named.Declaration in let hyps = List.rev (Environ.named_context env) in let map = function | LocalAssum (id, t) -> let t = EConstr.of_constr t in - Value.of_tuple [|Value.of_ident id; Value.of_option Value.of_constr None; Value.of_constr t|] + Value.of_tuple [|Value.of_ident id.binder_name; Value.of_option Value.of_constr None; Value.of_constr t|] | LocalDef (id, c, t) -> let c = EConstr.of_constr c in let t = EConstr.of_constr t in - Value.of_tuple [|Value.of_ident id; Value.of_option Value.of_constr (Some c); Value.of_constr t|] + Value.of_tuple [|Value.of_ident id.binder_name; Value.of_option Value.of_constr (Some c); Value.of_constr t|] in return (Value.of_list map hyps) end -- cgit v1.2.3 From 7c976ede65fbd5c6144e4cd58572c7c5a1229f73 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Wed, 20 Feb 2019 03:01:23 +0100 Subject: [coq] Fix OCaml warnings. In anticipation to https://github.com/coq/coq/pull/9605 , we fix all OCaml warnings. Fixes coq/ltac2#107 --- src/g_ltac2.mlg | 6 +++--- src/tac2core.ml | 12 +++++++----- src/tac2entries.ml | 20 ++++++++++---------- src/tac2env.ml | 2 +- src/tac2ffi.ml | 4 ++-- src/tac2intern.ml | 33 ++++++++++++++++----------------- src/tac2interp.ml | 2 +- src/tac2match.mli | 6 +++--- src/tac2print.ml | 4 ++-- src/tac2quote.ml | 12 ++++++------ 10 files changed, 51 insertions(+), 50 deletions(-) diff --git a/src/g_ltac2.mlg b/src/g_ltac2.mlg index 609e8a6c0e..7b058a339a 100644 --- a/src/g_ltac2.mlg +++ b/src/g_ltac2.mlg @@ -384,7 +384,7 @@ GRAMMAR EXTEND Gram ; END -(** Quotation scopes used by notations *) +(* Quotation scopes used by notations *) { @@ -866,8 +866,8 @@ end { -let pr_ltac2entry _ = mt () (** FIXME *) -let pr_ltac2expr _ = mt () (** FIXME *) +let pr_ltac2entry _ = mt () (* FIXME *) +let pr_ltac2expr _ = mt () (* FIXME *) } diff --git a/src/tac2core.ml b/src/tac2core.ml index b5ae446ed5..762a145318 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -404,6 +404,8 @@ let () = define1 "constr_kind" constr begin fun c -> Value.of_ext Value.val_projection p; Value.of_constr c; |] + | Int _ -> + assert false end end @@ -753,7 +755,7 @@ let () = define1 "hyp" ident begin fun id -> let mem = try ignore (Environ.lookup_named id env); true with Not_found -> false in if mem then return (Value.of_constr (EConstr.mkVar id)) else Tacticals.New.tclZEROMSG - (str "Hypothesis " ++ quote (Id.print id) ++ str " not found") (** FIXME: Do something more sensible *) + (str "Hypothesis " ++ quote (Id.print id) ++ str " not found") (* FIXME: Do something more sensible *) end end @@ -1082,7 +1084,7 @@ let () = let () = let intern self ist tac = - (** Prevent inner calls to Ltac2 values *) + (* Prevent inner calls to Ltac2 values *) let extra = Tac2intern.drop_ltac2_env ist.Genintern.extra in let ist = { ist with Genintern.extra } in let _, tac = Genintern.intern Ltac_plugin.Tacarg.wit_tactic ist tac in @@ -1113,7 +1115,7 @@ let () = let () = let open Ltac_plugin in let intern self ist tac = - (** Prevent inner calls to Ltac2 values *) + (* Prevent inner calls to Ltac2 values *) let extra = Tac2intern.drop_ltac2_env ist.Genintern.extra in let ist = { ist with Genintern.extra } in let _, tac = Genintern.intern Ltac_plugin.Tacarg.wit_tactic ist tac in @@ -1300,7 +1302,7 @@ end let () = add_scope "tactic" begin function | [] -> - (** Default to level 5 parsing *) + (* Default to level 5 parsing *) let scope = Extend.Aentryl (tac2expr, "5") in let act tac = tac in Tac2entries.ScopeRule (scope, act) @@ -1407,7 +1409,7 @@ let rec make_seq_rule = function let Seqrule (r, c) = make_seq_rule rem in let r = { norec_rule = Next (r.norec_rule, scope.any_symbol) } in let f = match tok with - | SexprStr _ -> None (** Leave out mere strings *) + | SexprStr _ -> None (* Leave out mere strings *) | _ -> Some f in Seqrule (r, CvCns (c, f)) diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 22025f0a8f..b7ce363957 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -111,7 +111,7 @@ let push_typedef visibility sp kn (_, def) = match def with | GTydDef _ -> Tac2env.push_type visibility sp kn | GTydAlg { galg_constructors = cstrs } -> - (** Register constructors *) + (* Register constructors *) let iter (c, _) = let spc = change_sp_label sp c in let knc = change_kn_label kn c in @@ -120,7 +120,7 @@ let push_typedef visibility sp kn (_, def) = match def with Tac2env.push_type visibility sp kn; List.iter iter cstrs | GTydRec fields -> - (** Register fields *) + (* Register fields *) let iter (c, _, _) = let spc = change_sp_label sp c in let knc = change_kn_label kn c in @@ -140,7 +140,7 @@ let define_typedef kn (params, def as qdef) = match def with | GTydDef _ -> Tac2env.define_type kn qdef | GTydAlg { galg_constructors = cstrs } -> - (** Define constructors *) + (* Define constructors *) let constant = ref 0 in let nonconstant = ref 0 in let iter (c, args) = @@ -157,7 +157,7 @@ let define_typedef kn (params, def as qdef) = match def with Tac2env.define_type kn qdef; List.iter iter cstrs | GTydRec fs -> - (** Define projections *) + (* Define projections *) let iter i (id, mut, t) = let knp = change_kn_label kn id in let proj = { @@ -297,7 +297,7 @@ let inline_rec_tactic tactics = let loc = pat.loc in (Id.Set.add id avoid, CAst.make ?loc id :: ans) in - (** Fresh variables to abstract over the function patterns *) + (* Fresh variables to abstract over the function patterns *) let _, vars = List.fold_left fold_var (avoid, []) pat in let map_body ({loc;v=id}, _, e) = CAst.(make ?loc @@ CPatVar (Name id)), e in let bnd = List.map map_body tactics in @@ -656,13 +656,13 @@ let inTac2Abbreviation : abbreviation -> obj = let register_notation ?(local = false) tkn lev body = match tkn, lev with | [SexprRec (_, {loc;v=Some id}, [])], None -> - (** Tactic abbreviation *) + (* Tactic abbreviation *) let () = check_lowercase CAst.(make ?loc id) in let body = Tac2intern.globalize Id.Set.empty body in let abbr = { abbr_body = body } in ignore (Lib.add_leaf id (inTac2Abbreviation abbr)) | _ -> - (** Check that the tokens make sense *) + (* Check that the tokens make sense *) let entries = List.map ParseToken.parse_token tkn in let fold accu tok = match tok with | TacTerm _ -> accu @@ -670,7 +670,7 @@ let register_notation ?(local = false) tkn lev body = match tkn, lev with | TacNonTerm (Anonymous, _) -> accu in let ids = List.fold_left fold Id.Set.empty entries in - (** Globalize so that names are absolute *) + (* Globalize so that names are absolute *) let body = Tac2intern.globalize ids body in let lev = match lev with Some _ -> lev | None -> Some 5 in let ext = { @@ -758,9 +758,9 @@ let perform_eval e = | Goal_select.SelectList l -> Proofview.tclFOCUSLIST l v | Goal_select.SelectId id -> Proofview.tclFOCUSID id v | Goal_select.SelectAll -> v - | Goal_select.SelectAlreadyFocused -> assert false (** TODO **) + | Goal_select.SelectAlreadyFocused -> assert false (* TODO **) in - (** HACK: the API doesn't allow to return a value *) + (* HACK: the API doesn't allow to return a value *) let ans = ref None in let tac = (v >>= fun r -> ans := Some r; Proofview.tclUNIT ()) in let (proof, _) = Proof.run_tactic (Global.env ()) tac proof in diff --git a/src/tac2env.ml b/src/tac2env.ml index 8198f92ff8..93ad57e97e 100644 --- a/src/tac2env.ml +++ b/src/tac2env.ml @@ -291,7 +291,7 @@ let is_constructor qid = let id = Id.to_string id in assert (String.length id > 0); match id with - | "true" | "false" -> true (** built-in constructors *) + | "true" | "false" -> true (* built-in constructors *) | _ -> match id.[0] with | 'A'..'Z' -> true diff --git a/src/tac2ffi.ml b/src/tac2ffi.ml index c271967bd6..e3127ab9df 100644 --- a/src/tac2ffi.ml +++ b/src/tac2ffi.ml @@ -349,12 +349,12 @@ let to_fun1 r0 r1 f = to_closure f let rec apply : type a. a arity -> a -> valexpr list -> valexpr Proofview.tactic = fun arity f args -> match args, arity with | [], arity -> Proofview.tclUNIT (ValCls (MLTactic (arity, f))) - (** A few hardcoded cases for efficiency *) + (* A few hardcoded cases for efficiency *) | [a0], OneAty -> f a0 | [a0; a1], AddAty OneAty -> f a0 a1 | [a0; a1; a2], AddAty (AddAty OneAty) -> f a0 a1 a2 | [a0; a1; a2; a3], AddAty (AddAty (AddAty OneAty)) -> f a0 a1 a2 a3 - (** Generic cases *) + (* Generic cases *) | a :: args, OneAty -> f a >>= fun f -> let MLTactic (arity, f) = to_closure f in diff --git a/src/tac2intern.ml b/src/tac2intern.ml index fe615853ce..de99fb167f 100644 --- a/src/tac2intern.ml +++ b/src/tac2intern.ml @@ -141,7 +141,7 @@ let empty_env () = { } let env_name env = - (** Generate names according to a provided environment *) + (* Generate names according to a provided environment *) let mk num = let base = num mod 26 in let rem = num / 26 in @@ -267,7 +267,6 @@ let fresh_reftype env (kn : KerName.t or_tuple) = (subst, t) (** First-order unification algorithm *) - let is_unfoldable kn = match snd (Tac2env.interp_type kn) with | GTydDef (Some _) -> true | GTydDef None | GTydAlg _ | GTydRec _ | GTydOpn -> false @@ -485,7 +484,7 @@ let check_elt_empty loc env t = match kind env t with let check_unit ?loc t = let env = empty_env () in - (** Should not matter, t should be closed. *) + (* Should not matter, t should be closed. *) let t = fresh_type_scheme env t in let maybe_unit = match kind env t with | GTypVar _ -> true @@ -618,7 +617,7 @@ let expand_pattern avoid bnd = let fold (avoid, bnd) (pat, t) = let na, expand = match pat.v with | CPatVar na -> - (** Don't expand variable patterns *) + (* Don't expand variable patterns *) na, None | _ -> let id = fresh_var avoid in @@ -691,7 +690,7 @@ let rec intern_rec env {loc;v=e} = match e with in let e = Tac2env.interp_alias kn in let map arg = - (** Thunk alias arguments *) + (* Thunk alias arguments *) let loc = arg.loc in let t_unit = CAst.make ?loc @@ CTypRef (AbsKn (Tuple 0), []) in let var = CAst.make ?loc @@ CPatCnv (CAst.make ?loc @@ CPatVar Anonymous, t_unit) in @@ -782,8 +781,8 @@ let rec intern_rec env {loc;v=e} = match e with intern_rec env e in let obj = interp_ml_object tag in - (** External objects do not have access to the named context because this is - not stable by dynamic semantics. *) + (* External objects do not have access to the named context because this is + not stable by dynamic semantics. *) let genv = Global.env_of_context Environ.empty_named_context_val in let ist = empty_glob_sign genv in let ist = { ist with extra = Store.set ist.extra ltac2_env env } in @@ -907,7 +906,7 @@ and intern_case env loc e pl = | CPatVar Anonymous -> let () = check_redundant_clause rem in let (br', brT) = intern_rec env br in - (** Fill all remaining branches *) + (* Fill all remaining branches *) let fill (ncst, narg) arity = if Int.equal arity 0 then let () = @@ -951,7 +950,7 @@ and intern_case env loc e pl = if not (Int.equal nids nargs) then error_nargs_mismatch ?loc knc nargs nids in let fold env id tpe = - (** Instantiate all arguments *) + (* Instantiate all arguments *) let subst n = GTypVar subst.(n) in let tpe = subst_type subst tpe in push_name id (monomorphic tpe) env @@ -1005,7 +1004,7 @@ and intern_case env loc e pl = let get = function | GPatVar na -> na | GPatRef _ -> - user_err ?loc (str "TODO: Unhandled match case") (** FIXME *) + user_err ?loc (str "TODO: Unhandled match case") (* FIXME *) in let loc = pat.loc in let knc = match knc with @@ -1024,7 +1023,7 @@ and intern_case env loc e pl = if not (Int.equal nids nargs) then error_nargs_mismatch ?loc knc nargs nids in let fold env id tpe = - (** Instantiate all arguments *) + (* Instantiate all arguments *) let subst n = GTypVar subst.(n) in let tpe = subst_type subst tpe in push_name id (monomorphic tpe) env @@ -1089,7 +1088,7 @@ and intern_record env loc fs = | _ -> assert false in let subst = Array.init params (fun _ -> fresh_id env) in - (** Set the answer [args] imperatively *) + (* Set the answer [args] imperatively *) let args = Array.make (List.length typdef) None in let iter (loc, pinfo, e) = if KerName.equal kn pinfo.pdata_type then @@ -1145,14 +1144,14 @@ let intern ~strict e = let intern_typedef self (ids, t) : glb_quant_typedef = let env = { (empty_env ()) with env_rec = self } in - (** Initialize type parameters *) + (* Initialize type parameters *) let map id = get_alias id env in let ids = List.map map ids in let count = ref (List.length ids) in let vars = ref UF.Map.empty in let iter n id = vars := UF.Map.add id (GTypVar n) !vars in let () = List.iteri iter ids in - (** Do not accept unbound type variables *) + (* Do not accept unbound type variables *) let env = { env with env_opn = false } in let intern t = let t = intern_type env t in @@ -1195,7 +1194,7 @@ let intern_open_type t = let check_subtype t1 t2 = let env = empty_env () in let t1 = fresh_type_scheme env t1 in - (** We build a substitution mimicking rigid variable by using dummy tuples *) + (* We build a substitution mimicking rigid variable by using dummy tuples *) let rigid i = GTypRef (Tuple (i + 1), []) in let (n, t2) = t2 in let subst = Array.init n rigid in @@ -1507,7 +1506,7 @@ let () = let intern ist tac = let env = match Genintern.Store.get ist.extra ltac2_env with | None -> - (** Only happens when Ltac2 is called from a constr or ltac1 quotation *) + (* Only happens when Ltac2 is called from a constr or ltac1 quotation *) let env = empty_env () in if !Ltac_plugin.Tacintern.strict_check then env else { env with env_str = false } @@ -1526,7 +1525,7 @@ let () = let intern ist (loc, id) = let env = match Genintern.Store.get ist.extra ltac2_env with | None -> - (** Only happens when Ltac2 is called from a constr or ltac1 quotation *) + (* Only happens when Ltac2 is called from a constr or ltac1 quotation *) let env = empty_env () in if !Ltac_plugin.Tacintern.strict_check then env else { env with env_str = false } diff --git a/src/tac2interp.ml b/src/tac2interp.ml index 6f158ac66e..b0f8083aeb 100644 --- a/src/tac2interp.ml +++ b/src/tac2interp.ml @@ -114,7 +114,7 @@ let rec interp (ist : environment) = function | Name id -> { env_ist = Id.Map.add id cls accu.env_ist } in let ist = List.fold_left fold ist fixs in - (** Hack to make a cycle imperatively in the environment *) + (* Hack to make a cycle imperatively in the environment *) let iter (_, e, _) = e.clos_env <- ist.env_ist in let () = List.iter iter fixs in interp ist e diff --git a/src/tac2match.mli b/src/tac2match.mli index 7cfa1ed25f..c82c40d238 100644 --- a/src/tac2match.mli +++ b/src/tac2match.mli @@ -28,6 +28,6 @@ val match_goal: constr -> rev:bool -> match_rule -> - ((Id.t * context option) list * (** List of hypotheses matching: name + context *) - context option * (** Context for conclusion *) - Ltac_pretype.patvar_map (** Pattern variable substitution *)) Proofview.tactic + ((Id.t * context option) list * (* List of hypotheses matching: name + context *) + context option * (* Context for conclusion *) + Ltac_pretype.patvar_map (* Pattern variable substitution *)) Proofview.tactic diff --git a/src/tac2print.ml b/src/tac2print.ml index 0b20cf9f58..f4cb290265 100644 --- a/src/tac2print.ml +++ b/src/tac2print.ml @@ -272,7 +272,7 @@ let pr_glbexpr_gen lvl c = paren (hov 0 (c ++ spc () ++ (pr_sequence (pr_glbexpr E0) cl))) | GTacExt (tag, arg) -> let tpe = interp_ml_object tag in - hov 0 (tpe.ml_print (Global.env ()) arg) (** FIXME *) + hov 0 (tpe.ml_print (Global.env ()) arg) (* FIXME *) | GTacPrm (prm, args) -> let args = match args with | [] -> mt () @@ -379,7 +379,7 @@ let rec pr_valexpr env sigma v t = match kind t with else match repr with | GTydDef None -> str "" | GTydDef (Some _) -> - (** Shouldn't happen thanks to kind *) + (* Shouldn't happen thanks to kind *) assert false | GTydAlg alg -> if Valexpr.is_int v then diff --git a/src/tac2quote.ml b/src/tac2quote.ml index 5a26e7465c..a98264745e 100644 --- a/src/tac2quote.ml +++ b/src/tac2quote.ml @@ -233,9 +233,9 @@ let abstract_vars loc vars tac = let na, tac = match def with | None -> (Anonymous, tac) | Some id0 -> - (** Trick: in order not to shadow a variable nor to choose an arbitrary - name, we reuse one which is going to be shadowed by the matched - variables anyways. *) + (* Trick: in order not to shadow a variable nor to choose an arbitrary + name, we reuse one which is going to be shadowed by the matched + variables anyways. *) let build_bindings (n, accu) na = match na with | Anonymous -> (n + 1, accu) | Name _ -> @@ -263,7 +263,7 @@ let of_conversion {loc;v=c} = match c with let vars = pattern_vars pat in let pat = of_option ?loc of_pattern (Some pat) in let c = of_constr c in - (** Order is critical here *) + (* Order is critical here *) let vars = List.map (fun id -> Name id) (Id.Set.elements vars) in let c = abstract_vars loc vars c in of_tuple [pat; c] @@ -388,7 +388,7 @@ let of_constr_matching {loc;v=m} = (knd, pat, na) in let vars = pattern_vars pat in - (** Order of elements is crucial here! *) + (* Order of elements is crucial here! *) let vars = Id.Set.elements vars in let vars = List.map (fun id -> Name id) vars in let e = abstract_vars loc vars tac in @@ -429,7 +429,7 @@ let of_goal_matching {loc;v=gm} = let hyps = List.map (fun ({CAst.v=na}, _, _, _) -> na) hyps_pats in let map (_, na, _, _) = na in let hctx = List.map map hyps_pats in - (** Order of elements is crucial here! *) + (* Order of elements is crucial here! *) let vars = Id.Set.elements vars in let subst = List.map (fun id -> Name id) vars in (r, hyps, hctx, subst, concl_ctx) -- cgit v1.2.3 From c5a6a100844803b2da370e4828655a9da377e624 Mon Sep 17 00:00:00 2001 From: Maxime Dénès Date: Fri, 15 Mar 2019 12:04:49 +0100 Subject: Adapt to changes in Coq's printers API --- src/tac2core.ml | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index e5499f0c73..15fd625650 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -1184,8 +1184,8 @@ let () = GlobEnv.register_constr_interp0 wit_ltac2_quotation interp let () = - let pr_raw id = Genprint.PrinterBasic mt in - let pr_glb id = Genprint.PrinterBasic (fun () -> str "$" ++ Id.print id) in + let pr_raw id = Genprint.PrinterBasic (fun _env _sigma -> mt ()) in + let pr_glb id = Genprint.PrinterBasic (fun _env _sigma -> str "$" ++ Id.print id) in let pr_top _ = Genprint.TopPrinterBasic mt in Genprint.register_print0 wit_ltac2_quotation pr_raw pr_glb pr_top @@ -1209,8 +1209,8 @@ let () = Geninterp.register_interp0 wit_ltac2 interp let () = - let pr_raw _ = Genprint.PrinterBasic mt in - let pr_glb e = Genprint.PrinterBasic (fun () -> Tac2print.pr_glbexpr e) in + let pr_raw _ = Genprint.PrinterBasic (fun _env _sigma -> mt ()) in + let pr_glb e = Genprint.PrinterBasic (fun _env _sigma -> Tac2print.pr_glbexpr e) in let pr_top _ = Genprint.TopPrinterBasic mt in Genprint.register_print0 wit_ltac2 pr_raw pr_glb pr_top -- cgit v1.2.3 From ea68d316725cd5012abf1012497e1c00e9bbb9d2 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Sun, 17 Feb 2019 02:54:17 +0100 Subject: [coq] Adapt to coq/coq#9129 "removal of imperative proof state" --- src/g_ltac2.mlg | 9 +++++---- src/tac2core.ml | 8 +++++--- src/tac2entries.ml | 28 +++++++++++++++------------- src/tac2entries.mli | 8 ++++++-- src/tac2tactics.ml | 1 + 5 files changed, 32 insertions(+), 22 deletions(-) diff --git a/src/g_ltac2.mlg b/src/g_ltac2.mlg index 7b058a339a..a404227d3d 100644 --- a/src/g_ltac2.mlg +++ b/src/g_ltac2.mlg @@ -890,8 +890,8 @@ let classify_ltac2 = function } VERNAC COMMAND EXTEND VernacDeclareTactic2Definition -| #[ local = locality ] [ "Ltac2" ltac2_entry(e) ] => { classify_ltac2 e } -> { - Tac2entries.register_struct ?local e +| #[ local = locality ] ![proof] [ "Ltac2" ltac2_entry(e) ] => { classify_ltac2 e } -> { + fun ~pstate -> Tac2entries.register_struct ?local ~pstate e; pstate } END @@ -914,10 +914,11 @@ open Vernacextend } VERNAC { tac2mode } EXTEND VernacLtac2 -| [ ltac2_expr(t) ltac_use_default(default) ] => +| ![proof] [ ltac2_expr(t) ltac_use_default(default) ] => { classify_as_proofstep } -> { (* let g = Option.default (Proof_global.get_default_goal_selector ()) g in *) - Tac2entries.call ~default t + fun ~pstate -> + Option.map (fun pstate -> Tac2entries.call ~pstate ~default t) pstate } END diff --git a/src/tac2core.ml b/src/tac2core.ml index 15fd625650..30beee58de 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -979,6 +979,7 @@ let constr_flags () = fail_evar = true; expand_evars = true; program_mode = false; + polymorphic = false; } let open_constr_no_classes_flags () = @@ -989,6 +990,7 @@ let open_constr_no_classes_flags () = fail_evar = false; expand_evars = true; program_mode = false; + polymorphic = false; } (** Embed all Ltac2 data into Values *) @@ -1164,17 +1166,17 @@ let () = (** Ltac2 in terms *) let () = - let interp ist env sigma concl tac = + let interp ist poly env sigma concl tac = let ist = Tac2interp.get_env ist in let tac = Proofview.tclIGNORE (Tac2interp.interp ist tac) in - let name, poly = Id.of_string "ltac2", false in + let name, poly = Id.of_string "ltac2", poly in let c, sigma = Pfedit.refine_by_tactic ~name ~poly env sigma concl tac in (EConstr.of_constr c, sigma) in GlobEnv.register_constr_interp0 wit_ltac2 interp let () = - let interp ist env sigma concl id = + let interp ist poly env sigma concl id = let ist = Tac2interp.get_env ist in let c = Id.Map.find id ist.env_ist in let c = Value.to_constr c in diff --git a/src/tac2entries.ml b/src/tac2entries.ml index b7ce363957..3073aad84f 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -739,19 +739,20 @@ let register_redefinition ?(local = false) qid e = } in Lib.add_anonymous_leaf (inTac2Redefinition def) -let perform_eval e = +let perform_eval ~pstate e = let open Proofview.Notations in let env = Global.env () in let (e, ty) = Tac2intern.intern ~strict:false e in let v = Tac2interp.interp Tac2interp.empty_environment e in let selector, proof = - try - Goal_select.get_default_goal_selector (), - Proof_global.give_me_the_proof () - with Proof_global.NoCurrentProof -> + match pstate with + | None -> let sigma = Evd.from_env env in let name, poly = Id.of_string "ltac2", false in Goal_select.SelectAll, Proof.start ~name ~poly sigma [] + | Some pstate -> + Goal_select.get_default_goal_selector (), + Proof_global.give_me_the_proof pstate in let v = match selector with | Goal_select.SelectNth i -> Proofview.tclFOCUS i i v @@ -773,13 +774,13 @@ let perform_eval e = (** Toplevel entries *) -let register_struct ?local str = match str with +let register_struct ?local ~pstate str = match str with | StrVal (mut, isrec, e) -> register_ltac ?local ~mut isrec e | StrTyp (isrec, t) -> register_type ?local isrec t | StrPrm (id, t, ml) -> register_primitive ?local id t ml | StrSyn (tok, lev, e) -> register_notation ?local tok lev e | StrMut (qid, e) -> register_redefinition ?local qid e -| StrRun e -> perform_eval e +| StrRun e -> perform_eval ~pstate e (** Toplevel exception *) @@ -860,8 +861,8 @@ let print_ltac qid = (** Calling tactics *) -let solve default tac = - let status = Proof_global.with_current_proof begin fun etac p -> +let solve ~pstate default tac = + let pstate, status = Proof_global.with_current_proof begin fun etac p -> let with_end_tac = if default then Some etac else None in let g = Goal_select.get_default_goal_selector () in let (p, status) = Pfedit.solve g None tac ?with_end_tac p in @@ -869,15 +870,16 @@ let solve default tac = go back to the top of the prooftree *) let p = Proof.maximal_unfocus Vernacentries.command_focus p in p, status - end in - if not status then Feedback.feedback Feedback.AddedAxiom + end pstate in + if not status then Feedback.feedback Feedback.AddedAxiom; + pstate -let call ~default e = +let call ~pstate ~default e = let loc = e.loc in let (e, t) = intern ~strict:false e in let () = check_unit ?loc t in let tac = Tac2interp.interp Tac2interp.empty_environment e in - solve default (Proofview.tclIGNORE tac) + solve ~pstate default (Proofview.tclIGNORE tac) (** Primitive algebraic types than can't be defined Coq-side *) diff --git a/src/tac2entries.mli b/src/tac2entries.mli index f97e35fec0..28b2120537 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -21,7 +21,11 @@ val register_type : ?local:bool -> rec_flag -> val register_primitive : ?local:bool -> Names.lident -> raw_typexpr -> ml_tactic_name -> unit -val register_struct : ?local:bool -> strexpr -> unit +val register_struct + : ?local:bool + -> pstate:Proof_global.t option + -> strexpr + -> unit val register_notation : ?local:bool -> sexpr list -> int option -> raw_tacexpr -> unit @@ -46,7 +50,7 @@ val print_ltac : Libnames.qualid -> unit (** {5 Eval loop} *) (** Evaluate a tactic expression in the current environment *) -val call : default:bool -> raw_tacexpr -> unit +val call : pstate:Proof_global.t -> default:bool -> raw_tacexpr -> Proof_global.t (** {5 Toplevel exceptions} *) diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index 059a1babd7..0c8522f495 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -24,6 +24,7 @@ let tactic_infer_flags with_evar = { Pretyping.fail_evar = not with_evar; Pretyping.expand_evars = true; Pretyping.program_mode = false; + Pretyping.polymorphic = false; } (** FIXME: export a better interface in Tactics *) -- cgit v1.2.3 From 9788d6d3b2799e5aeabaa4fbde9ae50d71f8b39b Mon Sep 17 00:00:00 2001 From: Vincent Laporte Date: Tue, 26 Mar 2019 08:46:48 +0000 Subject: Fix for https://github.com/coq/coq/pull/8984 --- src/tac2tactics.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/tac2tactics.ml b/src/tac2tactics.ml index 059a1babd7..5d9b14eab0 100644 --- a/src/tac2tactics.ml +++ b/src/tac2tactics.ml @@ -404,7 +404,7 @@ let eauto debug n p lems dbs = let typeclasses_eauto strategy depth dbs = let only_classes, dbs = match dbs with | None -> - true, [Hints.typeclasses_db] + true, [Class_tactics.typeclasses_db] | Some dbs -> let dbs = List.map Id.to_string dbs in false, dbs -- cgit v1.2.3 From 28d60e8f729ee6b66c9252c9766f3fe2d8d854cf Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Mon, 19 Nov 2018 20:10:59 +0100 Subject: [dune] Full Dune support. This add experimental support for building the full Ltac2 plugin with Dune, see tree at https://github.com/ejgallego/dune/tree/coq --- .gitignore | 3 +++ dune | 3 +++ dune-project | 3 +++ ltac2.opam | 0 src/dune | 6 +++--- theories/dune | 6 ++++++ 6 files changed, 18 insertions(+), 3 deletions(-) create mode 100644 dune create mode 100644 dune-project create mode 100644 ltac2.opam create mode 100644 theories/dune diff --git a/.gitignore b/.gitignore index 50ed772be3..00e15f8daa 100644 --- a/.gitignore +++ b/.gitignore @@ -13,3 +13,6 @@ Makefile.coq.conf *.a *.aux tests/*.log +*.install +_build +.merlin diff --git a/dune b/dune new file mode 100644 index 0000000000..5dbc4db66a --- /dev/null +++ b/dune @@ -0,0 +1,3 @@ +(env + (dev (flags :standard -rectypes)) + (release (flags :standard -rectypes))) diff --git a/dune-project b/dune-project new file mode 100644 index 0000000000..8154e999de --- /dev/null +++ b/dune-project @@ -0,0 +1,3 @@ +(lang dune 1.6) +(using coq 0.1) +(name ltac2) diff --git a/ltac2.opam b/ltac2.opam new file mode 100644 index 0000000000..e69de29bb2 diff --git a/src/dune b/src/dune index 4a018adb9a..332f3644b0 100644 --- a/src/dune +++ b/src/dune @@ -1,8 +1,8 @@ (library - (name ltac2) - (public_name coq.plugins.ltac2) + (name ltac2_plugin) + (public_name ltac2.plugin) (modules_without_implementation tac2expr tac2qexpr tac2types) - (flags :standard -w -50) + (flags :standard -warn-error -9-27-50) (libraries coq.plugins.firstorder)) (rule diff --git a/theories/dune b/theories/dune new file mode 100644 index 0000000000..1fe3ba28fe --- /dev/null +++ b/theories/dune @@ -0,0 +1,6 @@ +(coqlib + (name Ltac2) ; This determines the -R flag + (public_name ltac2.Ltac2) + (synopsis "Ltac 2 Plugin") + (libraries ltac2.plugin)) + -- cgit v1.2.3 From 93524ed2dfb3bbcc2006286954001039c95732cd Mon Sep 17 00:00:00 2001 From: Enrico Tassi Date: Wed, 13 Mar 2019 11:09:06 +0100 Subject: overlay for PR 9733 --- src/g_ltac2.mlg | 8 ++++---- src/tac2core.ml | 2 +- src/tac2entries.ml | 2 +- 3 files changed, 6 insertions(+), 6 deletions(-) diff --git a/src/g_ltac2.mlg b/src/g_ltac2.mlg index a404227d3d..2483ef459a 100644 --- a/src/g_ltac2.mlg +++ b/src/g_ltac2.mlg @@ -831,7 +831,7 @@ let open Tok in let (++) r s = Next (r, s) in let rules = [ Rule ( - Stop ++ Aentry test_dollar_ident ++ Atoken (KEYWORD "$") ++ Aentry Prim.ident, + Stop ++ Aentry test_dollar_ident ++ Atoken (pattern_for_KEYWORD "$") ++ Aentry Prim.ident, begin fun id _ _ loc -> let id = Loc.tag ~loc id in let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2_quotation) id in @@ -840,7 +840,7 @@ let rules = [ ); Rule ( - Stop ++ Aentry test_ampersand_ident ++ Atoken (KEYWORD "&") ++ Aentry Prim.ident, + Stop ++ Aentry test_ampersand_ident ++ Atoken (pattern_for_KEYWORD "&") ++ Aentry Prim.ident, begin fun id _ _ loc -> let tac = Tac2quote.of_exact_hyp ~loc (CAst.make ~loc id) in let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in @@ -849,8 +849,8 @@ let rules = [ ); Rule ( - Stop ++ Atoken (IDENT "ltac2") ++ Atoken (KEYWORD ":") ++ - Atoken (KEYWORD "(") ++ Aentry tac2expr ++ Atoken (KEYWORD ")"), + Stop ++ Atoken (pattern_for_IDENT "ltac2") ++ Atoken (pattern_for_KEYWORD ":") ++ + Atoken (pattern_for_KEYWORD "(") ++ Aentry tac2expr ++ Atoken (pattern_for_KEYWORD ")"), begin fun _ tac _ _ _ loc -> let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) diff --git a/src/tac2core.ml b/src/tac2core.ml index 30beee58de..0f85f56c22 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -1251,7 +1251,7 @@ open CAst let () = add_scope "keyword" begin function | [SexprStr {loc;v=s}] -> - let scope = Extend.Atoken (Tok.KEYWORD s) in + let scope = Extend.Atoken (Tok.pattern_for_KEYWORD s) in Tac2entries.ScopeRule (scope, (fun _ -> q_unit)) | arg -> scope_fail "keyword" arg end diff --git a/src/tac2entries.ml b/src/tac2entries.ml index 3073aad84f..bce58653e6 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -539,7 +539,7 @@ let parse_scope = function CErrors.user_err ?loc (str "Unknown scope" ++ spc () ++ Names.Id.print id) | SexprStr {v=str} -> let v_unit = CAst.make @@ CTacCst (AbsKn (Tuple 0)) in - ScopeRule (Extend.Atoken (Tok.IDENT str), (fun _ -> v_unit)) + ScopeRule (Extend.Atoken (Tok.pattern_for_IDENT str), (fun _ -> v_unit)) | tok -> let loc = loc_of_token tok in CErrors.user_err ?loc (str "Invalid parsing token") -- cgit v1.2.3 From 6a10b18d428178a9accd70a9717e153bb180868f Mon Sep 17 00:00:00 2001 From: Pierre Roux Date: Sun, 24 Mar 2019 22:53:38 +0100 Subject: [coq] Adapt to coq/coq#9815 --- src/g_ltac2.mlg | 8 ++++---- src/tac2core.ml | 38 ++++++++++++++------------------------ src/tac2entries.ml | 6 +++--- src/tac2entries.mli | 2 +- 4 files changed, 22 insertions(+), 32 deletions(-) diff --git a/src/g_ltac2.mlg b/src/g_ltac2.mlg index 2483ef459a..fcf5d59ec9 100644 --- a/src/g_ltac2.mlg +++ b/src/g_ltac2.mlg @@ -831,7 +831,7 @@ let open Tok in let (++) r s = Next (r, s) in let rules = [ Rule ( - Stop ++ Aentry test_dollar_ident ++ Atoken (pattern_for_KEYWORD "$") ++ Aentry Prim.ident, + Stop ++ Aentry test_dollar_ident ++ Atoken (PKEYWORD "$") ++ Aentry Prim.ident, begin fun id _ _ loc -> let id = Loc.tag ~loc id in let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2_quotation) id in @@ -840,7 +840,7 @@ let rules = [ ); Rule ( - Stop ++ Aentry test_ampersand_ident ++ Atoken (pattern_for_KEYWORD "&") ++ Aentry Prim.ident, + Stop ++ Aentry test_ampersand_ident ++ Atoken (PKEYWORD "&") ++ Aentry Prim.ident, begin fun id _ _ loc -> let tac = Tac2quote.of_exact_hyp ~loc (CAst.make ~loc id) in let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in @@ -849,8 +849,8 @@ let rules = [ ); Rule ( - Stop ++ Atoken (pattern_for_IDENT "ltac2") ++ Atoken (pattern_for_KEYWORD ":") ++ - Atoken (pattern_for_KEYWORD "(") ++ Aentry tac2expr ++ Atoken (pattern_for_KEYWORD ")"), + Stop ++ Atoken (PIDENT (Some "ltac2")) ++ Atoken (PKEYWORD ":") ++ + Atoken (PKEYWORD "(") ++ Aentry tac2expr ++ Atoken (PKEYWORD ")"), begin fun _ tac _ _ _ loc -> let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) diff --git a/src/tac2core.ml b/src/tac2core.ml index 0f85f56c22..a584933e00 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -1251,7 +1251,7 @@ open CAst let () = add_scope "keyword" begin function | [SexprStr {loc;v=s}] -> - let scope = Extend.Atoken (Tok.pattern_for_KEYWORD s) in + let scope = Extend.Atoken (Tok.PKEYWORD s) in Tac2entries.ScopeRule (scope, (fun _ -> q_unit)) | arg -> scope_fail "keyword" arg end @@ -1381,34 +1381,25 @@ let () = add_generic_scope "pattern" Pcoq.Constr.constr Tac2quote.wit_pattern open Extend exception SelfSymbol -type 'a any_symbol = { any_symbol : 'r. ('r, 'a) symbol } - let rec generalize_symbol : - type a s. (s, a) Extend.symbol -> a any_symbol = function -| Atoken tok -> - { any_symbol = Atoken tok } -| Alist1 e -> - let e = generalize_symbol e in - { any_symbol = Alist1 e.any_symbol } + type a tr s. (s, tr, a) Extend.symbol -> (s, Extend.norec, a) Extend.symbol = function +| Atoken tok -> Atoken tok +| Alist1 e -> Alist1 (generalize_symbol e) | Alist1sep (e, sep) -> let e = generalize_symbol e in let sep = generalize_symbol sep in - { any_symbol = Alist1sep (e.any_symbol, sep.any_symbol) } -| Alist0 e -> - let e = generalize_symbol e in - { any_symbol = Alist0 e.any_symbol } + Alist1sep (e, sep) +| Alist0 e -> Alist0 (generalize_symbol e) | Alist0sep (e, sep) -> let e = generalize_symbol e in let sep = generalize_symbol sep in - { any_symbol = Alist0sep (e.any_symbol, sep.any_symbol) } -| Aopt e -> - let e = generalize_symbol e in - { any_symbol = Aopt e.any_symbol } + Alist0sep (e, sep) +| Aopt e -> Aopt (generalize_symbol e) | Aself -> raise SelfSymbol | Anext -> raise SelfSymbol -| Aentry e -> { any_symbol = Aentry e } -| Aentryl (e, l) -> { any_symbol = Aentryl (e, l) } -| Arules r -> { any_symbol = Arules r } +| Aentry e -> Aentry e +| Aentryl (e, l) -> Aentryl (e, l) +| Arules r -> Arules r type _ converter = | CvNil : (Loc.t -> raw_tacexpr) converter @@ -1420,17 +1411,16 @@ let rec apply : type a. a converter -> raw_tacexpr list -> a = function | CvCns (c, Some f) -> fun accu x -> apply c (f x :: accu) type seqrule = -| Seqrule : ('act, Loc.t -> raw_tacexpr) norec_rule * 'act converter -> seqrule +| Seqrule : (Tac2expr.raw_tacexpr, Extend.norec, 'act, Loc.t -> raw_tacexpr) rule * 'act converter -> seqrule let rec make_seq_rule = function | [] -> - let r = { norec_rule = Stop } in - Seqrule (r, CvNil) + Seqrule (Stop, CvNil) | tok :: rem -> let Tac2entries.ScopeRule (scope, f) = Tac2entries.parse_scope tok in let scope = generalize_symbol scope in let Seqrule (r, c) = make_seq_rule rem in - let r = { norec_rule = Next (r.norec_rule, scope.any_symbol) } in + let r = NextNoRec (r, scope) in let f = match tok with | SexprStr _ -> None (* Leave out mere strings *) | _ -> Some f diff --git a/src/tac2entries.ml b/src/tac2entries.ml index bce58653e6..9fd01426de 100644 --- a/src/tac2entries.ml +++ b/src/tac2entries.ml @@ -514,7 +514,7 @@ type 'a token = | TacNonTerm of Name.t * 'a type scope_rule = -| ScopeRule : (raw_tacexpr, 'a) Extend.symbol * ('a -> raw_tacexpr) -> scope_rule +| ScopeRule : (raw_tacexpr, _, 'a) Extend.symbol * ('a -> raw_tacexpr) -> scope_rule type scope_interpretation = sexpr list -> scope_rule @@ -539,7 +539,7 @@ let parse_scope = function CErrors.user_err ?loc (str "Unknown scope" ++ spc () ++ Names.Id.print id) | SexprStr {v=str} -> let v_unit = CAst.make @@ CTacCst (AbsKn (Tuple 0)) in - ScopeRule (Extend.Atoken (Tok.pattern_for_IDENT str), (fun _ -> v_unit)) + ScopeRule (Extend.Atoken (Tok.PIDENT (Some str)), (fun _ -> v_unit)) | tok -> let loc = loc_of_token tok in CErrors.user_err ?loc (str "Invalid parsing token") @@ -567,7 +567,7 @@ type synext = { type krule = | KRule : - (raw_tacexpr, 'act, Loc.t -> raw_tacexpr) Extend.rule * + (raw_tacexpr, _, 'act, Loc.t -> raw_tacexpr) Extend.rule * ((Loc.t -> (Name.t * raw_tacexpr) list -> raw_tacexpr) -> 'act) -> krule let rec get_rule (tok : scope_rule token list) : krule = match tok with diff --git a/src/tac2entries.mli b/src/tac2entries.mli index 28b2120537..d493192bb3 100644 --- a/src/tac2entries.mli +++ b/src/tac2entries.mli @@ -33,7 +33,7 @@ val register_notation : ?local:bool -> sexpr list -> int option -> (** {5 Notations} *) type scope_rule = -| ScopeRule : (raw_tacexpr, 'a) Extend.symbol * ('a -> raw_tacexpr) -> scope_rule +| ScopeRule : (raw_tacexpr, _, 'a) Extend.symbol * ('a -> raw_tacexpr) -> scope_rule type scope_interpretation = sexpr list -> scope_rule -- cgit v1.2.3 From 80ad88ea8e2aab71c3dd0bf05b39776c61c93392 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Tue, 2 Apr 2019 12:36:27 +0200 Subject: [opam] Update file to newer format and build system. Using Dune in the OPAM file does allow to use some goodies such as `dune-release` etc... --- ltac2.opam | 18 ++++++++++++++++++ opam/descr | 1 - opam/opam | 17 ----------------- 3 files changed, 18 insertions(+), 18 deletions(-) delete mode 100644 opam/descr delete mode 100644 opam/opam diff --git a/ltac2.opam b/ltac2.opam index e69de29bb2..47ceb882b1 100644 --- a/ltac2.opam +++ b/ltac2.opam @@ -0,0 +1,18 @@ +synopsis: "A Tactic Language for Coq." +description: "A Tactic Language for Coq." +name: "coq-ltac2" +opam-version: "2.0" +maintainer: "Pierre-Marie Pédrot " +authors: "Pierre-Marie Pédrot " +homepage: "https://github.com/ppedrot/ltac2" +dev-repo: "https://github.com/ppedrot/ltac2.git" +bug-reports: "https://github.com/ppedrot/ltac2/issues" +license: "LGPL 2.1" +doc: "https://ppedrot.github.io/ltac2/doc" + +depends: [ + "coq" { = "dev" } + "dune" { build & >= "1.9.0" } +] + +build: [ "dune" "build" "-p" name "-j" jobs ] diff --git a/opam/descr b/opam/descr deleted file mode 100644 index 82463c4f45..0000000000 --- a/opam/descr +++ /dev/null @@ -1 +0,0 @@ -A tactic language for Coq. diff --git a/opam/opam b/opam/opam deleted file mode 100644 index e461b97942..0000000000 --- a/opam/opam +++ /dev/null @@ -1,17 +0,0 @@ -opam-version: "1.2" -name: "coq-ltac2" -version: "0.1" -maintainer: "Pierre-Marie Pédrot " -author: "Pierre-Marie Pédrot " -license: "LGPL 2.1" -homepage: "https://github.com/ppedrot/ltac2" -dev-repo: "https://github.com/ppedrot/ltac2.git" -bug-reports: "https://github.com/ppedrot/ltac2/issues" -build: [ - [make "COQBIN=\"\"" "-j%{jobs}%"] -] -install: [make "install"] -remove: [make "uninstall"] -depends: [ - "coq" { = "dev" } -] -- cgit v1.2.3 From 0d8a4ec0f7486f47b7cff4cda465e2bd10c163ac Mon Sep 17 00:00:00 2001 From: Pierre Roux Date: Sun, 27 Jan 2019 16:01:31 +0100 Subject: Adapt to coq/coq#8764 --- src/g_ltac2.mlg | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/g_ltac2.mlg b/src/g_ltac2.mlg index fcf5d59ec9..0071dbb088 100644 --- a/src/g_ltac2.mlg +++ b/src/g_ltac2.mlg @@ -44,7 +44,7 @@ let lk_ident n strm = match stream_nth n strm with | _ -> None let lk_int n strm = match stream_nth n strm with -| INT _ -> Some (n + 1) +| NUMERAL { NumTok.int = _; frac = ""; exp = "" } -> Some (n + 1) | _ -> None let lk_ident_or_anti = lk_ident <+> (lk_kw "$" >> lk_ident) -- cgit v1.2.3 From 5463cfbaaaab2e696c4bbeeeb38f03ca79d5949e Mon Sep 17 00:00:00 2001 From: Maxime Dénès Date: Tue, 9 Apr 2019 11:03:13 +0200 Subject: Adapt to Coq's PR #9909 --- src/tac2core.ml | 13 ++++++++++--- 1 file changed, 10 insertions(+), 3 deletions(-) diff --git a/src/tac2core.ml b/src/tac2core.ml index a584933e00..d7e7b91ee6 100644 --- a/src/tac2core.ml +++ b/src/tac2core.ml @@ -1030,9 +1030,10 @@ let () = let intern = intern_constr in let interp ist c = interp_constr (constr_flags ()) ist c in let print env c = str "constr:(" ++ Printer.pr_lglob_constr_env env c ++ str ")" in + let subst subst c = Detyping.subst_glob_constr (Global.env()) subst c in let obj = { ml_intern = intern; - ml_subst = Detyping.subst_glob_constr; + ml_subst = subst; ml_interp = interp; ml_print = print; } in @@ -1042,9 +1043,10 @@ let () = let intern = intern_constr in let interp ist c = interp_constr (open_constr_no_classes_flags ()) ist c in let print env c = str "open_constr:(" ++ Printer.pr_lglob_constr_env env c ++ str ")" in + let subst subst c = Detyping.subst_glob_constr (Global.env()) subst c in let obj = { ml_intern = intern; - ml_subst = Detyping.subst_glob_constr; + ml_subst = subst; ml_interp = interp; ml_print = print; } in @@ -1069,12 +1071,17 @@ let () = let _, pat = warn (fun () ->Constrintern.intern_constr_pattern env sigma ~as_type:false c) () in GlbVal pat, gtypref t_pattern in + let subst subst c = + let env = Global.env () in + let sigma = Evd.from_env env in + Patternops.subst_pattern env sigma subst c + in let print env pat = str "pattern:(" ++ Printer.pr_lconstr_pattern_env env Evd.empty pat ++ str ")" in let interp _ c = return (Value.of_pattern c) in let obj = { ml_intern = intern; ml_interp = interp; - ml_subst = Patternops.subst_pattern; + ml_subst = subst; ml_print = print; } in define_ml_object Tac2quote.wit_pattern obj -- cgit v1.2.3 From 66b6e83f4f4c32ad86333e13d65329be02c46048 Mon Sep 17 00:00:00 2001 From: Maxime Dénès Date: Thu, 25 Apr 2019 12:02:43 +0200 Subject: Prepare merge into Coq --- .gitignore | 18 - .travis.yml | 40 - LICENSE | 458 ---------- Makefile | 14 - README.md | 25 - _CoqProject | 51 -- doc/ltac2.md | 1036 ----------------------- dune | 3 - dune-project | 3 - ltac2.opam | 18 - src/dune | 11 - src/g_ltac2.mlg | 933 -------------------- src/ltac2_plugin.mlpack | 14 - src/tac2core.ml | 1446 ------------------------------- src/tac2core.mli | 30 - src/tac2dyn.ml | 27 - src/tac2dyn.mli | 34 - src/tac2entries.ml | 938 --------------------- src/tac2entries.mli | 93 -- src/tac2env.ml | 298 ------- src/tac2env.mli | 146 ---- src/tac2expr.mli | 190 ----- src/tac2extffi.ml | 40 - src/tac2extffi.mli | 16 - src/tac2ffi.ml | 382 --------- src/tac2ffi.mli | 189 ----- src/tac2intern.ml | 1545 ---------------------------------- src/tac2intern.mli | 46 - src/tac2interp.ml | 227 ----- src/tac2interp.mli | 37 - src/tac2match.ml | 232 ----- src/tac2match.mli | 33 - src/tac2print.ml | 488 ----------- src/tac2print.mli | 46 - src/tac2qexpr.mli | 173 ---- src/tac2quote.ml | 465 ---------- src/tac2quote.mli | 102 --- src/tac2stdlib.ml | 578 ------------- src/tac2stdlib.mli | 9 - src/tac2tactics.ml | 455 ---------- src/tac2tactics.mli | 124 --- src/tac2types.mli | 92 -- tests/Makefile | 16 - tests/compat.v | 58 -- tests/errors.v | 12 - tests/example1.v | 27 - tests/example2.v | 281 ------- tests/matching.v | 71 -- tests/quot.v | 26 - tests/rebind.v | 34 - tests/stuff/ltac2.v | 143 ---- tests/tacticals.v | 34 - tests/typing.v | 72 -- theories/Array.v | 14 - theories/Char.v | 12 - theories/Constr.v | 72 -- theories/Control.v | 76 -- theories/Env.v | 27 - theories/Fresh.v | 26 - theories/Ident.v | 17 - theories/Init.v | 69 -- theories/Int.v | 18 - theories/Ltac1.v | 36 - theories/Ltac2.v | 24 - theories/Message.v | 25 - theories/Notations.v | 568 ------------- theories/Pattern.v | 145 ---- theories/Std.v | 263 ------ theories/String.v | 14 - theories/dune | 6 - vendor/Ltac2/.gitignore | 18 + vendor/Ltac2/.travis.yml | 40 + vendor/Ltac2/LICENSE | 458 ++++++++++ vendor/Ltac2/Makefile | 14 + vendor/Ltac2/README.md | 25 + vendor/Ltac2/_CoqProject | 51 ++ vendor/Ltac2/doc/ltac2.md | 1036 +++++++++++++++++++++++ vendor/Ltac2/dune | 3 + vendor/Ltac2/dune-project | 3 + vendor/Ltac2/ltac2.opam | 18 + vendor/Ltac2/src/dune | 11 + vendor/Ltac2/src/g_ltac2.mlg | 933 ++++++++++++++++++++ vendor/Ltac2/src/ltac2_plugin.mlpack | 14 + vendor/Ltac2/src/tac2core.ml | 1446 +++++++++++++++++++++++++++++++ vendor/Ltac2/src/tac2core.mli | 30 + vendor/Ltac2/src/tac2dyn.ml | 27 + vendor/Ltac2/src/tac2dyn.mli | 34 + vendor/Ltac2/src/tac2entries.ml | 938 +++++++++++++++++++++ vendor/Ltac2/src/tac2entries.mli | 93 ++ vendor/Ltac2/src/tac2env.ml | 298 +++++++ vendor/Ltac2/src/tac2env.mli | 146 ++++ vendor/Ltac2/src/tac2expr.mli | 190 +++++ vendor/Ltac2/src/tac2extffi.ml | 40 + vendor/Ltac2/src/tac2extffi.mli | 16 + vendor/Ltac2/src/tac2ffi.ml | 382 +++++++++ vendor/Ltac2/src/tac2ffi.mli | 189 +++++ vendor/Ltac2/src/tac2intern.ml | 1545 ++++++++++++++++++++++++++++++++++ vendor/Ltac2/src/tac2intern.mli | 46 + vendor/Ltac2/src/tac2interp.ml | 227 +++++ vendor/Ltac2/src/tac2interp.mli | 37 + vendor/Ltac2/src/tac2match.ml | 232 +++++ vendor/Ltac2/src/tac2match.mli | 33 + vendor/Ltac2/src/tac2print.ml | 488 +++++++++++ vendor/Ltac2/src/tac2print.mli | 46 + vendor/Ltac2/src/tac2qexpr.mli | 173 ++++ vendor/Ltac2/src/tac2quote.ml | 465 ++++++++++ vendor/Ltac2/src/tac2quote.mli | 102 +++ vendor/Ltac2/src/tac2stdlib.ml | 578 +++++++++++++ vendor/Ltac2/src/tac2stdlib.mli | 9 + vendor/Ltac2/src/tac2tactics.ml | 455 ++++++++++ vendor/Ltac2/src/tac2tactics.mli | 124 +++ vendor/Ltac2/src/tac2types.mli | 92 ++ vendor/Ltac2/tests/Makefile | 16 + vendor/Ltac2/tests/compat.v | 58 ++ vendor/Ltac2/tests/errors.v | 12 + vendor/Ltac2/tests/example1.v | 27 + vendor/Ltac2/tests/example2.v | 281 +++++++ vendor/Ltac2/tests/matching.v | 71 ++ vendor/Ltac2/tests/quot.v | 26 + vendor/Ltac2/tests/rebind.v | 34 + vendor/Ltac2/tests/stuff/ltac2.v | 143 ++++ vendor/Ltac2/tests/tacticals.v | 34 + vendor/Ltac2/tests/typing.v | 72 ++ vendor/Ltac2/theories/Array.v | 14 + vendor/Ltac2/theories/Char.v | 12 + vendor/Ltac2/theories/Constr.v | 72 ++ vendor/Ltac2/theories/Control.v | 76 ++ vendor/Ltac2/theories/Env.v | 27 + vendor/Ltac2/theories/Fresh.v | 26 + vendor/Ltac2/theories/Ident.v | 17 + vendor/Ltac2/theories/Init.v | 69 ++ vendor/Ltac2/theories/Int.v | 18 + vendor/Ltac2/theories/Ltac1.v | 36 + vendor/Ltac2/theories/Ltac2.v | 24 + vendor/Ltac2/theories/Message.v | 25 + vendor/Ltac2/theories/Notations.v | 568 +++++++++++++ vendor/Ltac2/theories/Pattern.v | 145 ++++ vendor/Ltac2/theories/Std.v | 263 ++++++ vendor/Ltac2/theories/String.v | 14 + vendor/Ltac2/theories/dune | 6 + 140 files changed, 13291 insertions(+), 13291 deletions(-) delete mode 100644 .gitignore delete mode 100644 .travis.yml delete mode 100644 LICENSE delete mode 100644 Makefile delete mode 100644 README.md delete mode 100644 _CoqProject delete mode 100644 doc/ltac2.md delete mode 100644 dune delete mode 100644 dune-project delete mode 100644 ltac2.opam delete mode 100644 src/dune delete mode 100644 src/g_ltac2.mlg delete mode 100644 src/ltac2_plugin.mlpack delete mode 100644 src/tac2core.ml delete mode 100644 src/tac2core.mli delete mode 100644 src/tac2dyn.ml delete mode 100644 src/tac2dyn.mli delete mode 100644 src/tac2entries.ml delete mode 100644 src/tac2entries.mli delete mode 100644 src/tac2env.ml delete mode 100644 src/tac2env.mli delete mode 100644 src/tac2expr.mli delete mode 100644 src/tac2extffi.ml delete mode 100644 src/tac2extffi.mli delete mode 100644 src/tac2ffi.ml delete mode 100644 src/tac2ffi.mli delete mode 100644 src/tac2intern.ml delete mode 100644 src/tac2intern.mli delete mode 100644 src/tac2interp.ml delete mode 100644 src/tac2interp.mli delete mode 100644 src/tac2match.ml delete mode 100644 src/tac2match.mli delete mode 100644 src/tac2print.ml delete mode 100644 src/tac2print.mli delete mode 100644 src/tac2qexpr.mli delete mode 100644 src/tac2quote.ml delete mode 100644 src/tac2quote.mli delete mode 100644 src/tac2stdlib.ml delete mode 100644 src/tac2stdlib.mli delete mode 100644 src/tac2tactics.ml delete mode 100644 src/tac2tactics.mli delete mode 100644 src/tac2types.mli delete mode 100644 tests/Makefile delete mode 100644 tests/compat.v delete mode 100644 tests/errors.v delete mode 100644 tests/example1.v delete mode 100644 tests/example2.v delete mode 100644 tests/matching.v delete mode 100644 tests/quot.v delete mode 100644 tests/rebind.v delete mode 100644 tests/stuff/ltac2.v delete mode 100644 tests/tacticals.v delete mode 100644 tests/typing.v delete mode 100644 theories/Array.v delete mode 100644 theories/Char.v delete mode 100644 theories/Constr.v delete mode 100644 theories/Control.v delete mode 100644 theories/Env.v delete mode 100644 theories/Fresh.v delete mode 100644 theories/Ident.v delete mode 100644 theories/Init.v delete mode 100644 theories/Int.v delete mode 100644 theories/Ltac1.v delete mode 100644 theories/Ltac2.v delete mode 100644 theories/Message.v delete mode 100644 theories/Notations.v delete mode 100644 theories/Pattern.v delete mode 100644 theories/Std.v delete mode 100644 theories/String.v delete mode 100644 theories/dune create mode 100644 vendor/Ltac2/.gitignore create mode 100644 vendor/Ltac2/.travis.yml create mode 100644 vendor/Ltac2/LICENSE create mode 100644 vendor/Ltac2/Makefile create mode 100644 vendor/Ltac2/README.md create mode 100644 vendor/Ltac2/_CoqProject create mode 100644 vendor/Ltac2/doc/ltac2.md create mode 100644 vendor/Ltac2/dune create mode 100644 vendor/Ltac2/dune-project create mode 100644 vendor/Ltac2/ltac2.opam create mode 100644 vendor/Ltac2/src/dune create mode 100644 vendor/Ltac2/src/g_ltac2.mlg create mode 100644 vendor/Ltac2/src/ltac2_plugin.mlpack create mode 100644 vendor/Ltac2/src/tac2core.ml create mode 100644 vendor/Ltac2/src/tac2core.mli create mode 100644 vendor/Ltac2/src/tac2dyn.ml create mode 100644 vendor/Ltac2/src/tac2dyn.mli create mode 100644 vendor/Ltac2/src/tac2entries.ml create mode 100644 vendor/Ltac2/src/tac2entries.mli create mode 100644 vendor/Ltac2/src/tac2env.ml create mode 100644 vendor/Ltac2/src/tac2env.mli create mode 100644 vendor/Ltac2/src/tac2expr.mli create mode 100644 vendor/Ltac2/src/tac2extffi.ml create mode 100644 vendor/Ltac2/src/tac2extffi.mli create mode 100644 vendor/Ltac2/src/tac2ffi.ml create mode 100644 vendor/Ltac2/src/tac2ffi.mli create mode 100644 vendor/Ltac2/src/tac2intern.ml create mode 100644 vendor/Ltac2/src/tac2intern.mli create mode 100644 vendor/Ltac2/src/tac2interp.ml create mode 100644 vendor/Ltac2/src/tac2interp.mli create mode 100644 vendor/Ltac2/src/tac2match.ml create mode 100644 vendor/Ltac2/src/tac2match.mli create mode 100644 vendor/Ltac2/src/tac2print.ml create mode 100644 vendor/Ltac2/src/tac2print.mli create mode 100644 vendor/Ltac2/src/tac2qexpr.mli create mode 100644 vendor/Ltac2/src/tac2quote.ml create mode 100644 vendor/Ltac2/src/tac2quote.mli create mode 100644 vendor/Ltac2/src/tac2stdlib.ml create mode 100644 vendor/Ltac2/src/tac2stdlib.mli create mode 100644 vendor/Ltac2/src/tac2tactics.ml create mode 100644 vendor/Ltac2/src/tac2tactics.mli create mode 100644 vendor/Ltac2/src/tac2types.mli create mode 100644 vendor/Ltac2/tests/Makefile create mode 100644 vendor/Ltac2/tests/compat.v create mode 100644 vendor/Ltac2/tests/errors.v create mode 100644 vendor/Ltac2/tests/example1.v create mode 100644 vendor/Ltac2/tests/example2.v create mode 100644 vendor/Ltac2/tests/matching.v create mode 100644 vendor/Ltac2/tests/quot.v create mode 100644 vendor/Ltac2/tests/rebind.v create mode 100644 vendor/Ltac2/tests/stuff/ltac2.v create mode 100644 vendor/Ltac2/tests/tacticals.v create mode 100644 vendor/Ltac2/tests/typing.v create mode 100644 vendor/Ltac2/theories/Array.v create mode 100644 vendor/Ltac2/theories/Char.v create mode 100644 vendor/Ltac2/theories/Constr.v create mode 100644 vendor/Ltac2/theories/Control.v create mode 100644 vendor/Ltac2/theories/Env.v create mode 100644 vendor/Ltac2/theories/Fresh.v create mode 100644 vendor/Ltac2/theories/Ident.v create mode 100644 vendor/Ltac2/theories/Init.v create mode 100644 vendor/Ltac2/theories/Int.v create mode 100644 vendor/Ltac2/theories/Ltac1.v create mode 100644 vendor/Ltac2/theories/Ltac2.v create mode 100644 vendor/Ltac2/theories/Message.v create mode 100644 vendor/Ltac2/theories/Notations.v create mode 100644 vendor/Ltac2/theories/Pattern.v create mode 100644 vendor/Ltac2/theories/Std.v create mode 100644 vendor/Ltac2/theories/String.v create mode 100644 vendor/Ltac2/theories/dune diff --git a/.gitignore b/.gitignore deleted file mode 100644 index 00e15f8daa..0000000000 --- a/.gitignore +++ /dev/null @@ -1,18 +0,0 @@ -Makefile.coq -Makefile.coq.conf -*.glob -*.d -*.d.raw -*.vio -*.vo -*.cm* -*.annot -*.spit -*.spot -*.o -*.a -*.aux -tests/*.log -*.install -_build -.merlin diff --git a/.travis.yml b/.travis.yml deleted file mode 100644 index 2628abde45..0000000000 --- a/.travis.yml +++ /dev/null @@ -1,40 +0,0 @@ -dist: trusty -sudo: required -language: generic - -services: - - docker - -env: - global: - - NJOBS="2" - - CONTRIB_NAME="ltac2" - matrix: - - COQ_IMAGE="coqorg/coq:dev" - -install: | - # Prepare the COQ container - docker run -d -i --init --name=COQ -v ${TRAVIS_BUILD_DIR}:/home/coq/${CONTRIB_NAME} -w /home/coq/${CONTRIB_NAME} ${COQ_IMAGE} - docker exec COQ /bin/bash --login -c " - # This bash script is double-quoted to interpolate Travis CI env vars: - echo \"Build triggered by ${TRAVIS_EVENT_TYPE}\" - export PS4='+ \e[33;1m(\$0 @ line \$LINENO) \$\e[0m ' - set -ex # -e = exit on failure; -x = trace for debug - # opam update -y - # opam install -y -j ${NJOBS} coq-mathcomp-ssreflect - opam config list - opam repo list - opam list - " -script: -- echo -e "${ANSI_YELLOW}Building ${CONTRIB_NAME}...${ANSI_RESET}" && echo -en 'travis_fold:start:script\\r' -- | - docker exec COQ /bin/bash --login -c " - export PS4='+ \e[33;1m(\$0 @ line \$LINENO) \$\e[0m ' - set -ex - sudo chown -R coq:coq /home/coq/${CONTRIB_NAME} - make - make tests - " -- docker stop COQ # optional -- echo -en 'travis_fold:end:script\\r' diff --git a/LICENSE b/LICENSE deleted file mode 100644 index 27950e8d20..0000000000 --- a/LICENSE +++ /dev/null @@ -1,458 +0,0 @@ - GNU LESSER GENERAL PUBLIC LICENSE - Version 2.1, February 1999 - - Copyright (C) 1991, 1999 Free Software Foundation, Inc. - 59 Temple Place, Suite 330, Boston, MA 02111-1307 USA - Everyone is permitted to copy and distribute verbatim copies - of this license document, but changing it is not allowed. - -[This is the first released version of the Lesser GPL. It also counts - as the successor of the GNU Library Public License, version 2, hence - the version number 2.1.] - - Preamble - - The licenses for most software are designed to take away your -freedom to share and change it. By contrast, the GNU General Public -Licenses are intended to guarantee your freedom to share and change -free software--to make sure the software is free for all its users. - - This license, the Lesser General Public License, applies to some -specially designated software packages--typically libraries--of the -Free Software Foundation and other authors who decide to use it. You -can use it too, but we suggest you first think carefully about whether -this license or the ordinary General Public License is the better -strategy to use in any particular case, based on the explanations below. - - When we speak of free software, we are referring to freedom of use, -not price. Our General Public Licenses are designed to make sure that -you have the freedom to distribute copies of free software (and charge -for this service if you wish); that you receive source code or can get -it if you want it; that you can change the software and use pieces of -it in new free programs; and that you are informed that you can do -these things. - - To protect your rights, we need to make restrictions that forbid -distributors to deny you these rights or to ask you to surrender these -rights. These restrictions translate to certain responsibilities for -you if you distribute copies of the library or if you modify it. - - For example, if you distribute copies of the library, whether gratis -or for a fee, you must give the recipients all the rights that we gave -you. You must make sure that they, too, receive or can get the source -code. If you link other code with the library, you must provide -complete object files to the recipients, so that they can relink them -with the library after making changes to the library and recompiling -it. And you must show them these terms so they know their rights. - - We protect your rights with a two-step method: (1) we copyright the -library, and (2) we offer you this license, which gives you legal -permission to copy, distribute and/or modify the library. - - To protect each distributor, we want to make it very clear that -there is no warranty for the free library. Also, if the library is -modified by someone else and passed on, the recipients should know -that what they have is not the original version, so that the original -author's reputation will not be affected by problems that might be -introduced by others. - - Finally, software patents pose a constant threat to the existence of -any free program. We wish to make sure that a company cannot -effectively restrict the users of a free program by obtaining a -restrictive license from a patent holder. Therefore, we insist that -any patent license obtained for a version of the library must be -consistent with the full freedom of use specified in this license. - - Most GNU software, including some libraries, is covered by the -ordinary GNU General Public License. This license, the GNU Lesser -General Public License, applies to certain designated libraries, and -is quite different from the ordinary General Public License. We use -this license for certain libraries in order to permit linking those -libraries into non-free programs. - - When a program is linked with a library, whether statically or using -a shared library, the combination of the two is legally speaking a -combined work, a derivative of the original library. The ordinary -General Public License therefore permits such linking only if the -entire combination fits its criteria of freedom. The Lesser General -Public License permits more lax criteria for linking other code with -the library. - - We call this license the "Lesser" General Public License because it -does Less to protect the user's freedom than the ordinary General -Public License. It also provides other free software developers Less -of an advantage over competing non-free programs. These disadvantages -are the reason we use the ordinary General Public License for many -libraries. However, the Lesser license provides advantages in certain -special circumstances. - - For example, on rare occasions, there may be a special need to -encourage the widest possible use of a certain library, so that it becomes -a de-facto standard. To achieve this, non-free programs must be -allowed to use the library. A more frequent case is that a free -library does the same job as widely used non-free libraries. In this -case, there is little to gain by limiting the free library to free -software only, so we use the Lesser General Public License. - - In other cases, permission to use a particular library in non-free -programs enables a greater number of people to use a large body of -free software. For example, permission to use the GNU C Library in -non-free programs enables many more people to use the whole GNU -operating system, as well as its variant, the GNU/Linux operating -system. - - Although the Lesser General Public License is Less protective of the -users' freedom, it does ensure that the user of a program that is -linked with the Library has the freedom and the wherewithal to run -that program using a modified version of the Library. - - The precise terms and conditions for copying, distribution and -modification follow. Pay close attention to the difference between a -"work based on the library" and a "work that uses the library". The -former contains code derived from the library, whereas the latter must -be combined with the library in order to run. - - GNU LESSER GENERAL PUBLIC LICENSE - TERMS AND CONDITIONS FOR COPYING, DISTRIBUTION AND MODIFICATION - - 0. This License Agreement applies to any software library or other -program which contains a notice placed by the copyright holder or -other authorized party saying it may be distributed under the terms of -this Lesser General Public License (also called "this License"). -Each licensee is addressed as "you". - - A "library" means a collection of software functions and/or data -prepared so as to be conveniently linked with application programs -(which use some of those functions and data) to form executables. - - The "Library", below, refers to any such software library or work -which has been distributed under these terms. A "work based on the -Library" means either the Library or any derivative work under -copyright law: that is to say, a work containing the Library or a -portion of it, either verbatim or with modifications and/or translated -straightforwardly into another language. (Hereinafter, translation is -included without limitation in the term "modification".) - - "Source code" for a work means the preferred form of the work for -making modifications to it. For a library, complete source code means -all the source code for all modules it contains, plus any associated -interface definition files, plus the scripts used to control compilation -and installation of the library. - - Activities other than copying, distribution and modification are not -covered by this License; they are outside its scope. The act of -running a program using the Library is not restricted, and output from -such a program is covered only if its contents constitute a work based -on the Library (independent of the use of the Library in a tool for -writing it). Whether that is true depends on what the Library does -and what the program that uses the Library does. - - 1. You may copy and distribute verbatim copies of the Library's -complete source code as you receive it, in any medium, provided that -you conspicuously and appropriately publish on each copy an -appropriate copyright notice and disclaimer of warranty; keep intact -all the notices that refer to this License and to the absence of any -warranty; and distribute a copy of this License along with the -Library. - - You may charge a fee for the physical act of transferring a copy, -and you may at your option offer warranty protection in exchange for a -fee. - - 2. You may modify your copy or copies of the Library or any portion -of it, thus forming a work based on the Library, and copy and -distribute such modifications or work under the terms of Section 1 -above, provided that you also meet all of these conditions: - - a) The modified work must itself be a software library. - - b) You must cause the files modified to carry prominent notices - stating that you changed the files and the date of any change. - - c) You must cause the whole of the work to be licensed at no - charge to all third parties under the terms of this License. - - d) If a facility in the modified Library refers to a function or a - table of data to be supplied by an application program that uses - the facility, other than as an argument passed when the facility - is invoked, then you must make a good faith effort to ensure that, - in the event an application does not supply such function or - table, the facility still operates, and performs whatever part of - its purpose remains meaningful. - - (For example, a function in a library to compute square roots has - a purpose that is entirely well-defined independent of the - application. Therefore, Subsection 2d requires that any - application-supplied function or table used by this function must - be optional: if the application does not supply it, the square - root function must still compute square roots.) - -These requirements apply to the modified work as a whole. If -identifiable sections of that work are not derived from the Library, -and can be reasonably considered independent and separate works in -themselves, then this License, and its terms, do not apply to those -sections when you distribute them as separate works. But when you -distribute the same sections as part of a whole which is a work based -on the Library, the distribution of the whole must be on the terms of -this License, whose permissions for other licensees extend to the -entire whole, and thus to each and every part regardless of who wrote -it. - -Thus, it is not the intent of this section to claim rights or contest -your rights to work written entirely by you; rather, the intent is to -exercise the right to control the distribution of derivative or -collective works based on the Library. - -In addition, mere aggregation of another work not based on the Library -with the Library (or with a work based on the Library) on a volume of -a storage or distribution medium does not bring the other work under -the scope of this License. - - 3. You may opt to apply the terms of the ordinary GNU General Public -License instead of this License to a given copy of the Library. To do -this, you must alter all the notices that refer to this License, so -that they refer to the ordinary GNU General Public License, version 2, -instead of to this License. (If a newer version than version 2 of the -ordinary GNU General Public License has appeared, then you can specify -that version instead if you wish.) Do not make any other change in -these notices. - - Once this change is made in a given copy, it is irreversible for -that copy, so the ordinary GNU General Public License applies to all -subsequent copies and derivative works made from that copy. - - This option is useful when you wish to copy part of the code of -the Library into a program that is not a library. - - 4. You may copy and distribute the Library (or a portion or -derivative of it, under Section 2) in object code or executable form -under the terms of Sections 1 and 2 above provided that you accompany -it with the complete corresponding machine-readable source code, which -must be distributed under the terms of Sections 1 and 2 above on a -medium customarily used for software interchange. - - If distribution of object code is made by offering access to copy -from a designated place, then offering equivalent access to copy the -source code from the same place satisfies the requirement to -distribute the source code, even though third parties are not -compelled to copy the source along with the object code. - - 5. A program that contains no derivative of any portion of the -Library, but is designed to work with the Library by being compiled or -linked with it, is called a "work that uses the Library". Such a -work, in isolation, is not a derivative work of the Library, and -therefore falls outside the scope of this License. - - However, linking a "work that uses the Library" with the Library -creates an executable that is a derivative of the Library (because it -contains portions of the Library), rather than a "work that uses the -library". The executable is therefore covered by this License. -Section 6 states terms for distribution of such executables. - - When a "work that uses the Library" uses material from a header file -that is part of the Library, the object code for the work may be a -derivative work of the Library even though the source code is not. -Whether this is true is especially significant if the work can be -linked without the Library, or if the work is itself a library. The -threshold for this to be true is not precisely defined by law. - - If such an object file uses only numerical parameters, data -structure layouts and accessors, and small macros and small inline -functions (ten lines or less in length), then the use of the object -file is unrestricted, regardless of whether it is legally a derivative -work. (Executables containing this object code plus portions of the -Library will still fall under Section 6.) - - Otherwise, if the work is a derivative of the Library, you may -distribute the object code for the work under the terms of Section 6. -Any executables containing that work also fall under Section 6, -whether or not they are linked directly with the Library itself. - - 6. As an exception to the Sections above, you may also combine or -link a "work that uses the Library" with the Library to produce a -work containing portions of the Library, and distribute that work -under terms of your choice, provided that the terms permit -modification of the work for the customer's own use and reverse -engineering for debugging such modifications. - - You must give prominent notice with each copy of the work that the -Library is used in it and that the Library and its use are covered by -this License. You must supply a copy of this License. If the work -during execution displays copyright notices, you must include the -copyright notice for the Library among them, as well as a reference -directing the user to the copy of this License. Also, you must do one -of these things: - - a) Accompany the work with the complete corresponding - machine-readable source code for the Library including whatever - changes were used in the work (which must be distributed under - Sections 1 and 2 above); and, if the work is an executable linked - with the Library, with the complete machine-readable "work that - uses the Library", as object code and/or source code, so that the - user can modify the Library and then relink to produce a modified - executable containing the modified Library. (It is understood - that the user who changes the contents of definitions files in the - Library will not necessarily be able to recompile the application - to use the modified definitions.) - - b) Use a suitable shared library mechanism for linking with the - Library. A suitable mechanism is one that (1) uses at run time a - copy of the library already present on the user's computer system, - rather than copying library functions into the executable, and (2) - will operate properly with a modified version of the library, if - the user installs one, as long as the modified version is - interface-compatible with the version that the work was made with. - - c) Accompany the work with a written offer, valid for at - least three years, to give the same user the materials - specified in Subsection 6a, above, for a charge no more - than the cost of performing this distribution. - - d) If distribution of the work is made by offering access to copy - from a designated place, offer equivalent access to copy the above - specified materials from the same place. - - e) Verify that the user has already received a copy of these - materials or that you have already sent this user a copy. - - For an executable, the required form of the "work that uses the -Library" must include any data and utility programs needed for -reproducing the executable from it. However, as a special exception, -the materials to be distributed need not include anything that is -normally distributed (in either source or binary form) with the major -components (compiler, kernel, and so on) of the operating system on -which the executable runs, unless that component itself accompanies -the executable. - - It may happen that this requirement contradicts the license -restrictions of other proprietary libraries that do not normally -accompany the operating system. Such a contradiction means you cannot -use both them and the Library together in an executable that you -distribute. - - 7. You may place library facilities that are a work based on the -Library side-by-side in a single library together with other library -facilities not covered by this License, and distribute such a combined -library, provided that the separate distribution of the work based on -the Library and of the other library facilities is otherwise -permitted, and provided that you do these two things: - - a) Accompany the combined library with a copy of the same work - based on the Library, uncombined with any other library - facilities. This must be distributed under the terms of the - Sections above. - - b) Give prominent notice with the combined library of the fact - that part of it is a work based on the Library, and explaining - where to find the accompanying uncombined form of the same work. - - 8. You may not copy, modify, sublicense, link with, or distribute -the Library except as expressly provided under this License. Any -attempt otherwise to copy, modify, sublicense, link with, or -distribute the Library is void, and will automatically terminate your -rights under this License. However, parties who have received copies, -or rights, from you under this License will not have their licenses -terminated so long as such parties remain in full compliance. - - 9. You are not required to accept this License, since you have not -signed it. However, nothing else grants you permission to modify or -distribute the Library or its derivative works. These actions are -prohibited by law if you do not accept this License. Therefore, by -modifying or distributing the Library (or any work based on the -Library), you indicate your acceptance of this License to do so, and -all its terms and conditions for copying, distributing or modifying -the Library or works based on it. - - 10. Each time you redistribute the Library (or any work based on the -Library), the recipient automatically receives a license from the -original licensor to copy, distribute, link with or modify the Library -subject to these terms and conditions. You may not impose any further -restrictions on the recipients' exercise of the rights granted herein. -You are not responsible for enforcing compliance by third parties with -this License. - - 11. If, as a consequence of a court judgment or allegation of patent -infringement or for any other reason (not limited to patent issues), -conditions are imposed on you (whether by court order, agreement or -otherwise) that contradict the conditions of this License, they do not -excuse you from the conditions of this License. If you cannot -distribute so as to satisfy simultaneously your obligations under this -License and any other pertinent obligations, then as a consequence you -may not distribute the Library at all. For example, if a patent -license would not permit royalty-free redistribution of the Library by -all those who receive copies directly or indirectly through you, then -the only way you could satisfy both it and this License would be to -refrain entirely from distribution of the Library. - -If any portion of this section is held invalid or unenforceable under any -particular circumstance, the balance of the section is intended to apply, -and the section as a whole is intended to apply in other circumstances. - -It is not the purpose of this section to induce you to infringe any -patents or other property right claims or to contest validity of any -such claims; this section has the sole purpose of protecting the -integrity of the free software distribution system which is -implemented by public license practices. Many people have made -generous contributions to the wide range of software distributed -through that system in reliance on consistent application of that -system; it is up to the author/donor to decide if he or she is willing -to distribute software through any other system and a licensee cannot -impose that choice. - -This section is intended to make thoroughly clear what is believed to -be a consequence of the rest of this License. - - 12. If the distribution and/or use of the Library is restricted in -certain countries either by patents or by copyrighted interfaces, the -original copyright holder who places the Library under this License may add -an explicit geographical distribution limitation excluding those countries, -so that distribution is permitted only in or among countries not thus -excluded. In such case, this License incorporates the limitation as if -written in the body of this License. - - 13. The Free Software Foundation may publish revised and/or new -versions of the Lesser General Public License from time to time. -Such new versions will be similar in spirit to the present version, -but may differ in detail to address new problems or concerns. - -Each version is given a distinguishing version number. If the Library -specifies a version number of this License which applies to it and -"any later version", you have the option of following the terms and -conditions either of that version or of any later version published by -the Free Software Foundation. If the Library does not specify a -license version number, you may choose any version ever published by -the Free Software Foundation. - - 14. If you wish to incorporate parts of the Library into other free -programs whose distribution conditions are incompatible with these, -write to the author to ask for permission. For software which is -copyrighted by the Free Software Foundation, write to the Free -Software Foundation; we sometimes make exceptions for this. Our -decision will be guided by the two goals of preserving the free status -of all derivatives of our free software and of promoting the sharing -and reuse of software generally. - - NO WARRANTY - - 15. BECAUSE THE LIBRARY IS LICENSED FREE OF CHARGE, THERE IS NO -WARRANTY FOR THE LIBRARY, TO THE EXTENT PERMITTED BY APPLICABLE LAW. -EXCEPT WHEN OTHERWISE STATED IN WRITING THE COPYRIGHT HOLDERS AND/OR -OTHER PARTIES PROVIDE THE LIBRARY "AS IS" WITHOUT WARRANTY OF ANY -KIND, EITHER EXPRESSED OR IMPLIED, INCLUDING, BUT NOT LIMITED TO, THE -IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR -PURPOSE. THE ENTIRE RISK AS TO THE QUALITY AND PERFORMANCE OF THE -LIBRARY IS WITH YOU. SHOULD THE LIBRARY PROVE DEFECTIVE, YOU ASSUME -THE COST OF ALL NECESSARY SERVICING, REPAIR OR CORRECTION. - - 16. IN NO EVENT UNLESS REQUIRED BY APPLICABLE LAW OR AGREED TO IN -WRITING WILL ANY COPYRIGHT HOLDER, OR ANY OTHER PARTY WHO MAY MODIFY -AND/OR REDISTRIBUTE THE LIBRARY AS PERMITTED ABOVE, BE LIABLE TO YOU -FOR DAMAGES, INCLUDING ANY GENERAL, SPECIAL, INCIDENTAL OR -CONSEQUENTIAL DAMAGES ARISING OUT OF THE USE OR INABILITY TO USE THE -LIBRARY (INCLUDING BUT NOT LIMITED TO LOSS OF DATA OR DATA BEING -RENDERED INACCURATE OR LOSSES SUSTAINED BY YOU OR THIRD PARTIES OR A -FAILURE OF THE LIBRARY TO OPERATE WITH ANY OTHER SOFTWARE), EVEN IF -SUCH HOLDER OR OTHER PARTY HAS BEEN ADVISED OF THE POSSIBILITY OF SUCH -DAMAGES. - - END OF TERMS AND CONDITIONS diff --git a/Makefile b/Makefile deleted file mode 100644 index e0e197650d..0000000000 --- a/Makefile +++ /dev/null @@ -1,14 +0,0 @@ -ifeq "$(COQBIN)" "" - COQBIN=$(dir $(shell which coqtop))/ -endif - -%: Makefile.coq - -Makefile.coq: _CoqProject - $(COQBIN)coq_makefile -f _CoqProject -o Makefile.coq - -tests: all - @$(MAKE) -C tests -s clean - @$(MAKE) -C tests -s all - --include Makefile.coq diff --git a/README.md b/README.md deleted file mode 100644 index d49dd88076..0000000000 --- a/README.md +++ /dev/null @@ -1,25 +0,0 @@ -[![Build Status](https://travis-ci.org/ppedrot/ltac2.svg?branch=master)](https://travis-ci.org/ppedrot/ltac2) - -Overview -======== - -This is a standalone version of the Ltac2 plugin. Ltac2 is an attempt at -providing the Coq users with a tactic language that is more robust and more -expressive than the venerable Ltac language. - -Status -======== - -It is mostly a toy to experiment for now, and the implementation is quite -bug-ridden. Don't mistake this for a final product! - -Installation -============ - -This should compile with Coq master, assuming the `COQBIN` variable is set -correctly. Standard procedures for `coq_makefile`-generated plugins apply. - -Demo -==== - -Horrible test-files are provided in the `tests` folder. Not for kids. diff --git a/_CoqProject b/_CoqProject deleted file mode 100644 index dda5a8001a..0000000000 --- a/_CoqProject +++ /dev/null @@ -1,51 +0,0 @@ --R theories/ Ltac2 --I src/ - -src/tac2dyn.ml -src/tac2dyn.mli -src/tac2expr.mli -src/tac2types.mli -src/tac2env.ml -src/tac2env.mli -src/tac2print.ml -src/tac2print.mli -src/tac2intern.ml -src/tac2intern.mli -src/tac2interp.ml -src/tac2interp.mli -src/tac2entries.ml -src/tac2entries.mli -src/tac2ffi.ml -src/tac2ffi.mli -src/tac2qexpr.mli -src/tac2quote.ml -src/tac2quote.mli -src/tac2match.ml -src/tac2match.mli -src/tac2core.ml -src/tac2core.mli -src/tac2extffi.ml -src/tac2extffi.mli -src/tac2tactics.ml -src/tac2tactics.mli -src/tac2stdlib.ml -src/tac2stdlib.mli -src/g_ltac2.mlg -src/ltac2_plugin.mlpack - -theories/Init.v -theories/Int.v -theories/Char.v -theories/String.v -theories/Ident.v -theories/Array.v -theories/Control.v -theories/Message.v -theories/Constr.v -theories/Pattern.v -theories/Fresh.v -theories/Std.v -theories/Env.v -theories/Notations.v -theories/Ltac1.v -theories/Ltac2.v diff --git a/doc/ltac2.md b/doc/ltac2.md deleted file mode 100644 index b217cb08e6..0000000000 --- a/doc/ltac2.md +++ /dev/null @@ -1,1036 +0,0 @@ -# Summary - -The Ltac tactic language is probably one of the ingredients of the success of -Coq, yet it is at the same time its Achilles' heel. Indeed, Ltac: - -- has nothing like intended semantics -- is very non-uniform due to organic growth -- lacks expressivity and requires programming-by-hacking -- is slow -- is error-prone and fragile -- has an intricate implementation - -This has a lot of terrible consequences, most notably the fact that it is never -clear whether some observed behaviour is a bug or a proper one. - -Following the need of users that start developing huge projects relying -critically on Ltac, we believe that we should offer a proper modern language -that features at least the following: - -- at least informal, predictable semantics -- a typing system -- standard programming facilities (i.e. datatypes) - -This document describes the implementation of such a language. The -implementation of Ltac as of Coq 8.7 will be referred to as Ltac1. - -# Contents - - -**Table of Contents** - -- [Summary](#summary) -- [Contents](#contents) -- [General design](#general-design) -- [ML component](#ml-component) - - [Overview](#overview) - - [Type Syntax](#type-syntax) - - [Type declarations](#type-declarations) - - [Term Syntax](#term-syntax) - - [Ltac Definitions](#ltac-definitions) - - [Reduction](#reduction) - - [Typing](#typing) - - [Effects](#effects) - - [Standard IO](#standard-io) - - [Fatal errors](#fatal-errors) - - [Backtrack](#backtrack) - - [Goals](#goals) -- [Meta-programming](#meta-programming) - - [Overview](#overview-1) - - [Generic Syntax for Quotations](#generic-syntax-for-quotations) - - [Built-in quotations](#built-in-quotations) - - [Strict vs. non-strict mode](#strict-vs-non-strict-mode) - - [Term Antiquotations](#term-antiquotations) - - [Syntax](#syntax) - - [Semantics](#semantics) - - [Static semantics](#static-semantics) - - [Dynamic semantics](#dynamic-semantics) - - [Trivial Term Antiquotations](#trivial-term-antiquotations) - - [Match over terms](#match-over-terms) - - [Match over goals](#match-over-goals) -- [Notations](#notations) - - [Scopes](#scopes) - - [Notations](#notations-1) - - [Abbreviations](#abbreviations) -- [Evaluation](#evaluation) -- [Debug](#debug) -- [Compatibility layer with Ltac1](#compatibility-layer-with-ltac1) - - [Ltac1 from Ltac2](#ltac1-from-ltac2) - - [Ltac2 from Ltac1](#ltac2-from-ltac1) -- [Transition from Ltac1](#transition-from-ltac1) - - [Syntax changes](#syntax-changes) - - [Tactic delay](#tactic-delay) - - [Variable binding](#variable-binding) - - [In Ltac expressions](#in-ltac-expressions) - - [In quotations](#in-quotations) - - [Exception catching](#exception-catching) -- [TODO](#todo) - - - - -# General design - -There are various alternatives to Ltac1, such that Mtac or Rtac for instance. -While those alternatives can be quite distinct from Ltac1, we designed -Ltac2 to be closest as reasonably possible to Ltac1, while fixing the -aforementioned defects. - -In particular, Ltac2 is: -- a member of the ML family of languages, i.e. - * a call-by-value functional language - * with effects - * together with Hindley-Milner type system -- a language featuring meta-programming facilities for the manipulation of - Coq-side terms -- a language featuring notation facilities to help writing palatable scripts - -We describe more in details each point in the remainder of this document. - -# ML component - -## Overview - -Ltac2 is a member of the ML family of languages, in the sense that it is an -effectful call-by-value functional language, with static typing à la -Hindley-Milner. It is commonly accepted that ML constitutes a sweet spot in PL -design, as it is relatively expressive while not being either too lax -(contrarily to dynamic typing) nor too strict (contrarily to say, dependent -types). - -The main goal of Ltac2 is to serve as a meta-language for Coq. As such, it -naturally fits in the ML lineage, just as the historical ML was designed as -the tactic language for the LCF prover. It can also be seen as a general-purpose -language, by simply forgetting about the Coq-specific features. - -Sticking to a standard ML type system can be considered somewhat weak for a -meta-language designed to manipulate Coq terms. In particular, there is no -way to statically guarantee that a Coq term resulting from an Ltac2 -computation will be well-typed. This is actually a design choice, motivated -by retro-compatibility with Ltac1. Instead, well-typedness is deferred to -dynamic checks, allowing many primitive functions to fail whenever they are -provided with an ill-typed term. - -The language is naturally effectful as it manipulates the global state of the -proof engine. This allows to think of proof-modifying primitives as effects -in a straightforward way. Semantically, proof manipulation lives in a monad, -which allows to ensure that Ltac2 satisfies the same equations as a generic ML -with unspecified effects would do, e.g. function reduction is substitution -by a value. - -## Type Syntax - -At the level of terms, we simply elaborate on Ltac1 syntax, which is quite -close to e.g. the one of OCaml. Types follow the simply-typed syntax of OCaml. - -``` -TYPE := -| "(" TYPE₀ "," ... "," TYPEₙ ")" TYPECONST -| "(" TYPE₀ "*" ... "*" TYPEₙ ")" -| TYPE₁ "->" TYPE₂ -| TYPEVAR - -TYPECONST := ( MODPATH "." )* LIDENT - -TYPEVAR := "'" LIDENT - -TYPEPARAMS := "(" TYPEVAR₀ "," ... "," TYPEVARₙ ")" -``` - -The set of base types can be extended thanks to the usual ML type -declarations such as algebraic datatypes and records. - -Built-in types include: -- `int`, machine integers (size not specified, in practice inherited from OCaml) -- `string`, mutable strings -- `'a array`, mutable arrays -- `exn`, exceptions -- `constr`, kernel-side terms -- `pattern`, term patterns -- `ident`, well-formed identifiers - -## Type declarations - -One can define new types by the following commands. - -``` -VERNAC ::= -| "Ltac2" "Type" TYPEPARAMS LIDENT -| "Ltac2" "Type" RECFLAG TYPEPARAMS LIDENT ":=" TYPEDEF - -RECFLAG := ( "rec" ) -``` - -The first command defines an abstract type. It has no use for the end user and -is dedicated to types representing data coming from the OCaml world. - -The second command defines a type with a manifest. There are four possible -kinds of such definitions: alias, variant, record and open variant types. - -``` -TYPEDEF := -| TYPE -| "[" CONSTRUCTORDEF₀ "|" ... "|" CONSTRUCTORDEFₙ "]" -| "{" FIELDDEF₀ ";" ... ";" FIELDDEFₙ "}" -| "[" ".." "]" - -CONSTRUCTORDEF := -| IDENT ( "(" TYPE₀ "," ... "," TYPE₀ ")" ) - -FIELDDEF := -| MUTFLAG IDENT ":" TYPE - -MUTFLAG := ( "mutable" ) -``` - -Aliases are just a name for a given type expression and are transparently -unfoldable to it. They cannot be recursive. - -Variants are sum types defined by constructors and eliminated by -pattern-matching. They can be recursive, but the `RECFLAG` must be explicitly -set. Pattern-maching must be exhaustive. - -Records are product types with named fields and eliminated by projection. -Likewise they can be recursive if the `RECFLAG` is set. - -Open variants are a special kind of variant types whose constructors are not -statically defined, but can instead be extended dynamically. A typical example -is the standard `exn` type. Pattern-matching must always include a catch-all -clause. They can be extended by the following command. - -``` -VERNAC ::= -| "Ltac2" "Type" TYPEPARAMS QUALID ":=" "[" CONSTRUCTORDEF "]" -``` - -## Term Syntax - -The syntax of the functional fragment is very close to the one of Ltac1, except -that it adds a true pattern-matching feature, as well as a few standard -constructions from ML. - -``` -VAR := LIDENT - -QUALID := ( MODPATH "." )* LIDENT - -CONSTRUCTOR := UIDENT - -TERM := -| QUALID -| CONSTRUCTOR TERM₀ ... TERMₙ -| TERM TERM₀ ... TERMₙ -| "fun" VAR "=>" TERM -| "let" VAR ":=" TERM "in" TERM -| "let" "rec" VAR ":=" TERM "in" TERM -| "match" TERM "with" BRANCH* "end" -| INT -| STRING -| "[|" TERM₀ ";" ... ";" TERMₙ "|]" -| "(" TERM₀ "," ... "," TERMₙ ")" -| "{" FIELD+ "}" -| TERM "." "(" QUALID ")" -| TERM₁ "." "(" QUALID ")" ":=" TERM₂ -| "["; TERM₀ ";" ... ";" TERMₙ "]" -| TERM₁ "::" TERM₂ -| ... - -BRANCH := -| PATTERN "=>" TERM - -PATTERN := -| VAR -| "_" -| "(" PATTERN₀ "," ... "," PATTERNₙ ")" -| CONSTRUCTOR PATTERN₀ ... PATTERNₙ -| "[" "]" -| PATTERN₁ "::" PATTERN₂ - -FIELD := -| QUALID ":=" TERM - -``` - -In practice, there is some additional syntactic sugar that allows e.g. to -bind a variable and match on it at the same time, in the usual ML style. - -There is a dedicated syntax for list and array litterals. - -Limitations: for now, deep pattern matching is not implemented yet. - -## Ltac Definitions - -One can define a new global Ltac2 value using the following syntax. -``` -VERNAC ::= -| "Ltac2" MUTFLAG RECFLAG LIDENT ":=" TERM -``` - -For semantic reasons, the body of the Ltac2 definition must be a syntactical -value, i.e. a function, a constant or a pure constructor recursively applied to -values. - -If the `RECFLAG` is set, the tactic is expanded into a recursive binding. - -If the `MUTFLAG` is set, the definition can be redefined at a later stage. This -can be performed through the following command. - -``` -VERNAC ::= -| "Ltac2" "Set" QUALID ":=" TERM -``` - -Mutable definitions act like dynamic binding, i.e. at runtime, the last defined -value for this entry is chosen. This is useful for global flags and the like. - -## Reduction - -We use the usual ML call-by-value reduction, with an otherwise unspecified -evaluation order. This is a design choice making it compatible with OCaml, -if ever we implement native compilation. The expected equations are as follows. -``` -(fun x => t) V ≡ t{x := V} (βv) - -let x := V in t ≡ t{x := V} (let) - -match C V₀ ... Vₙ with ... | C x₀ ... xₙ => t | ... end ≡ t {xᵢ := Vᵢ} (ι) - -(t any term, V values, C constructor) -``` - -Note that call-by-value reduction is already a departure from Ltac1 which uses -heuristics to decide when evaluating an expression. For instance, the following -expressions do not evaluate the same way in Ltac1. - -``` -foo (idtac; let x := 0 in bar) - -foo (let x := 0 in bar) -``` - -Instead of relying on the `idtac` hack, we would now require an explicit thunk -not to compute the argument, and `foo` would have e.g. type -`(unit -> unit) -> unit`. - -``` -foo (fun () => let x := 0 in bar) -``` - -## Typing - -Typing is strict and follows Hindley-Milner system. We will not implement the -current hackish subtyping semantics, and one will have to resort to conversion -functions. See notations though to make things more palatable. - -In this setting, all usual argument-free tactics have type `unit -> unit`, but -one can return as well a value of type `t` thanks to terms of type `unit -> t`, -or take additional arguments. - -## Effects - -Regarding effects, nothing involved here, except that instead of using the -standard IO monad as the ambient effectful world, Ltac2 is going to use the -tactic monad. - -Note that the order of evaluation of application is *not* specified and is -implementation-dependent, as in OCaml. - -We recall that the `Proofview.tactic` monad is essentially a IO monad together -with backtracking state representing the proof state. - -Intuitively a thunk of type `unit -> 'a` can do the following: -- It can perform non-backtracking IO like printing and setting mutable variables -- It can fail in a non-recoverable way -- It can use first-class backtrack. The proper way to figure that is that we - morally have the following isomorphism: - `(unit -> 'a) ~ (unit -> exn + ('a * (exn -> 'a)))` i.e. thunks can produce a - lazy list of results where each tail is waiting for a continuation exception. -- It can access a backtracking proof state, made out amongst other things of - the current evar assignation and the list of goals under focus. - -We describe more thoroughly the various effects existing in Ltac2 hereafter. - -### Standard IO - -The Ltac2 language features non-backtracking IO, notably mutable data and -printing operations. - -Mutable fields of records can be modified using the set syntax. Likewise, -built-in types like `string` and `array` feature imperative assignment. See -modules `String` and `Array` respectively. - -A few printing primitives are provided in the `Message` module, allowing to -display information to the user. - -### Fatal errors - -The Ltac2 language provides non-backtracking exceptions through the -following primitive in module `Control`. - -``` -val throw : exn -> 'a -``` - -Contrarily to backtracking exceptions from the next section, this kind of error -is never caught by backtracking primitives, that is, throwing an exception -destroys the stack. This is materialized by the following equation, where `E` -is an evaluation context. - -``` -E[throw e] ≡ throw e - -(e value) -``` - -There is currently no way to catch such an exception and it is a design choice. -There might be at some future point a way to catch it in a brutal way, -destroying all backtrack and return values. - -### Backtrack - -In Ltac2, we have the following backtracking primitives, defined in the -`Control` module. - -``` -Ltac2 Type 'a result := [ Val ('a) | Err (exn) ]. - -val zero : exn -> 'a -val plus : (unit -> 'a) -> (exn -> 'a) -> 'a -val case : (unit -> 'a) -> ('a * (exn -> 'a)) result -``` - -If one sees thunks as lazy lists, then `zero` is the empty list and `plus` is -list concatenation, while `case` is pattern-matching. - -The backtracking is first-class, i.e. one can write -`plus (fun () => "x") (fun _ => "y") : string` producing a backtracking string. - -These operations are expected to satisfy a few equations, most notably that they -form a monoid compatible with sequentialization. - -``` -plus t zero ≡ t () -plus (fun () => zero e) f ≡ f e -plus (plus t f) g ≡ plus t (fun e => plus (f e) g) - -case (fun () => zero e) ≡ Err e -case (fun () => plus (fun () => t) f) ≡ Val t f - -let x := zero e in u ≡ zero e -let x := plus t f in u ≡ plus (fun () => let x := t in u) (fun e => let x := f e in u) - -(t, u, f, g, e values) -``` - -### Goals - -A goal is given by the data of its conclusion and hypotheses, i.e. it can be -represented as `[Γ ⊢ A]`. - -The tactic monad naturally operates over the whole proofview, which may -represent several goals, including none. Thus, there is no such thing as -*the current goal*. Goals are naturally ordered, though. - -It is natural to do the same in Ltac2, but we must provide a way to get access -to a given goal. This is the role of the `enter` primitive, that applies a -tactic to each currently focused goal in turn. - -``` -val enter : (unit -> unit) -> unit -``` - -It is guaranteed that when evaluating `enter f`, `f` is called with exactly one -goal under focus. Note that `f` may be called several times, or never, depending -on the number of goals under focus before the call to `enter`. - -Accessing the goal data is then implicit in the Ltac2 primitives, and may panic -if the invariants are not respected. The two essential functions for observing -goals are given below. - -``` -val hyp : ident -> constr -val goal : unit -> constr -``` - -The two above functions panic if there is not exactly one goal under focus. -In addition, `hyp` may also fail if there is no hypothesis with the -corresponding name. - -# Meta-programming - -## Overview - -One of the horrendous implementation issues of Ltac is the fact that it is -never clear whether an object refers to the object world or the meta-world. -This is an incredible source of slowness, as the interpretation must be -aware of bound variables and must use heuristics to decide whether a variable -is a proper one or referring to something in the Ltac context. - -Likewise, in Ltac1, constr parsing is implicit, so that `foo 0` is -not `foo` applied to the Ltac integer expression `0` (Ltac does have a -non-first-class notion of integers), but rather the Coq term `Datatypes.O`. - -We should stop doing that! We need to mark when quoting and unquoting, although -we need to do that in a short and elegant way so as not to be too cumbersome -to the user. - -## Generic Syntax for Quotations - -In general, quotations can be introduced in term by the following syntax, where -`QUOTENTRY` is some parsing entry. -``` -TERM ::= -| QUOTNAME ":" "(" QUOTENTRY ")" - -QUOTNAME := IDENT -``` - -### Built-in quotations - -The current implementation recognizes the following built-in quotations: -- "ident", which parses identifiers (type `Init.ident`). -- "constr", which parses Coq terms and produces an-evar free term at runtime - (type `Init.constr`). -- "open_constr", which parses Coq terms and produces a term potentially with - holes at runtime (type `Init.constr` as well). -- "pattern", which parses Coq patterns and produces a pattern used for term - matching (type `Init.pattern`). -- "reference", which parses either a `QUALID` or `"&" IDENT`. Qualified names - are globalized at internalization into the corresponding global reference, - while `&id` is turned into `Std.VarRef id`. This produces at runtime a - `Std.reference`. - -The following syntactic sugar is provided for two common cases. -- `@id` is the same as ident:(id) -- `'t` is the same as open_constr:(t) - -### Strict vs. non-strict mode - -Depending on the context, quotations producing terms (i.e. `constr` or -`open_constr`) are not internalized in the same way. There are two possible -modes, respectively called the *strict* and the *non-strict* mode. - -- In strict mode, all simple identifiers appearing in a term quotation are -required to be resolvable statically. That is, they must be the short name of -a declaration which is defined globally, excluding section variables and -hypotheses. If this doesn't hold, internalization will fail. To work around -this error, one has to specifically use the `&` notation. -- In non-strict mode, any simple identifier appearing in a term quotation which -is not bound in the global context is turned into a dynamic reference to a -hypothesis. That is to say, internalization will succeed, but the evaluation -of the term at runtime will fail if there is no such variable in the dynamic -context. - -Strict mode is enforced by default, e.g. for all Ltac2 definitions. Non-strict -mode is only set when evaluating Ltac2 snippets in interactive proof mode. The -rationale is that it is cumbersome to explicitly add `&` interactively, while it -is expected that global tactics enforce more invariants on their code. - -## Term Antiquotations - -### Syntax - -One can also insert Ltac2 code into Coq term, similarly to what was possible in -Ltac1. - -``` -COQCONSTR ::= -| "ltac2" ":" "(" TERM ")" -``` - -Antiquoted terms are expected to have type `unit`, as they are only evaluated -for their side-effects. - -### Semantics - -Interpretation of a quoted Coq term is done in two phases, internalization and -evaluation. - -- Internalization is part of the static semantics, i.e. it is done at Ltac2 - typing time. -- Evaluation is part of the dynamic semantics, i.e. it is done when - a term gets effectively computed by Ltac2. - -Remark that typing of Coq terms is a *dynamic* process occuring at Ltac2 -evaluation time, and not at Ltac2 typing time. - -#### Static semantics - -During internalization, Coq variables are resolved and antiquotations are -type-checked as Ltac2 terms, effectively producing a `glob_constr` in Coq -implementation terminology. Note that although it went through the -type-checking of **Ltac2**, the resulting term has not been fully computed and -is potentially ill-typed as a runtime **Coq** term. - -``` -Ltac2 Definition myconstr () := constr:(nat -> 0). -// Valid with type `unit -> constr`, but will fail at runtime. -``` - -Term antiquotations are type-checked in the enclosing Ltac2 typing context -of the corresponding term expression. For instance, the following will -type-check. - -``` -let x := '0 in constr:(1 + ltac2:(exact x)) -// type constr -``` - -Beware that the typing environment of typing of antiquotations is **not** -expanded by the Coq binders from the term. Namely, it means that the following -Ltac2 expression will **not** type-check. -``` -constr:(fun x : nat => ltac2:(exact x)) -// Error: Unbound variable 'x' -``` - -There is a simple reason for that, which is that the following expression would -not make sense in general. -``` -constr:(fun x : nat => ltac2:(clear @x; exact x)) -``` -Indeed, a hypothesis can suddenly disappear from the runtime context if some -other tactic pulls the rug from under you. - -Rather, the tactic writer has to resort to the **dynamic** goal environment, -and must write instead explicitly that she is accessing a hypothesis, typically -as follows. -``` -constr:(fun x : nat => ltac2:(exact (hyp @x))) -``` - -This pattern is so common that we provide dedicated Ltac2 and Coq term notations -for it. - -- `&x` as an Ltac2 expression expands to `hyp @x`. -- `&x` as a Coq constr expression expands to - `ltac2:(Control.refine (fun () => hyp @x))`. - -#### Dynamic semantics - -During evaluation, a quoted term is fully evaluated to a kernel term, and is -in particular type-checked in the current environment. - -Evaluation of a quoted term goes as follows. -- The quoted term is first evaluated by the pretyper. -- Antiquotations are then evaluated in a context where there is exactly one goal -under focus, with the hypotheses coming from the current environment extended -with the bound variables of the term, and the resulting term is fed into the -quoted term. - -Relative orders of evaluation of antiquotations and quoted term are not -specified. - -For instance, in the following example, `tac` will be evaluated in a context -with exactly one goal under focus, whose last hypothesis is `H : nat`. The -whole expression will thus evaluate to the term `fun H : nat => nat`. -``` -let tac () := hyp @H in constr:(fun H : nat => ltac2:(tac ())) -``` - -Many standard tactics perform type-checking of their argument before going -further. It is your duty to ensure that terms are well-typed when calling -such tactics. Failure to do so will result in non-recoverable exceptions. - -## Trivial Term Antiquotations - -It is possible to refer to a variable of type `constr` in the Ltac2 environment -through a specific syntax consistent with the antiquotations presented in -the notation section. - -``` -COQCONSTR ::= -| "$" LIDENT -``` - -In a Coq term, writing `$x` is semantically equivalent to -`ltac2:(Control.refine (fun () => x))`, up to re-typechecking. It allows to -insert in a concise way an Ltac2 variable of type `constr` into a Coq term. - -## Match over terms - -Ltac2 features a construction similar to Ltac1 `match` over terms, although -in a less hard-wired way. - -``` -TERM ::= -| "match!" TERM "with" CONSTR-MATCHING* "end" -| "lazy_match!" TERM "with" CONSTR-MATCHING* "end" -| "multi_match!" TERM "with" CONSTR-MATCHING*"end" - -CONSTR-MATCHING := -| "|" CONSTR-PATTERN "=>" TERM - -CONSTR-PATTERN := -| CONSTR -| "context" LIDENT? "[" CONSTR "]" -``` - -This construction is not primitive and is desugared at parsing time into -calls to term matching functions from the `Pattern` module. Internally, it is -implemented thanks to a specific scope accepting the `CONSTR-MATCHING` syntax. - -Variables from the `CONSTR-PATTERN` are statically bound in the body of the branch, to -values of type `constr` for the variables from the `CONSTR` pattern and to a -value of type `Pattern.context` for the variable `LIDENT`. - -Note that contrarily to Ltac, only lowercase identifiers are valid as Ltac2 -bindings, so that there will be a syntax error if one of the bound variables -starts with an uppercase character. - -The semantics of this construction is otherwise the same as the corresponding -one from Ltac1, except that it requires the goal to be focused. - -## Match over goals - -Similarly, there is a way to match over goals in an elegant way, which is -just a notation desugared at parsing time. - -``` -TERM ::= -| "match!" MATCH-ORDER? "goal" "with" GOAL-MATCHING* "end" -| "lazy_match!" MATCH-ORDER? "goal" "with" GOAL-MATCHING* "end" -| "multi_match!" MATCH-ORDER? "goal" "with" GOAL-MATCHING*"end" - -GOAL-MATCHING := -| "|" "[" HYP-MATCHING* "|-" CONSTR-PATTERN "]" "=>" TERM - -HYP-MATCHING := -| LIDENT ":" CONSTR-PATTERN - -MATCH-ORDER := -| "reverse" -``` - -Variables from `HYP-MATCHING` and `CONSTR-PATTERN` are bound in the body of the -branch. Their types are: -- `constr` for pattern variables appearing in a `CONSTR` -- `Pattern.context` for variables binding a context -- `ident` for variables binding a hypothesis name. - -The same identifier caveat as in the case of matching over constr applies, and -this features has the same semantics as in Ltac1. In particular, a `reverse` -flag can be specified to match hypotheses from the more recently introduced to -the least recently introduced one. - -# Notations - -Notations are the crux of the usability of Ltac1. We should be able to recover -a feeling similar to the old implementation by using and abusing notations. - -## Scopes - -A scope is a name given to a grammar entry used to produce some Ltac2 expression -at parsing time. Scopes are described using a form of S-expression. - -``` -SCOPE := -| STRING -| INT -| LIDENT ( "(" SCOPE₀ "," ... "," SCOPEₙ ")" ) -``` - -A few scopes contain antiquotation features. For sake of uniformity, all -antiquotations are introduced by the syntax `"$" VAR`. - -The following scopes are built-in. -- constr: - + parses `c = COQCONSTR` and produces `constr:(c)` -- ident: - + parses `id = IDENT` and produces `ident:(id)` - + parses `"$" (x = IDENT)` and produces the variable `x` -- list0(*scope*): - + if *scope* parses `ENTRY`, parses ̀`(x₀, ..., xₙ = ENTRY*)` and produces - `[x₀; ...; xₙ]`. -- list0(*scope*, sep = STRING): - + if *scope* parses `ENTRY`, parses `(x₀ = ENTRY, "sep", ..., "sep", xₙ = ENTRY)` - and produces `[x₀; ...; xₙ]`. -- list1: same as list0 (with or without separator) but parses `ENTRY+` instead - of `ENTRY*`. -- opt(*scope*) - + if *scope* parses `ENTRY`, parses `ENTRY?` and produces either `None` or - `Some x` where `x` is the parsed expression. -- self: - + parses a Ltac2 expression at the current level and return it as is. -- next: - + parses a Ltac2 expression at the next level and return it as is. -- tactic(n = INT): - + parses a Ltac2 expression at the provided level *n* and return it as is. -- thunk(*scope*): - + parses the same as *scope*, and if *e* is the parsed expression, returns - `fun () => e`. -- STRING: - + parses the corresponding string as a CAMLP5 IDENT and returns `()`. -- keyword(s = STRING): - + parses the string *s* as a keyword and returns `()`. -- terminal(s = STRING): - + parses the string *s* as a keyword, if it is already a - keyword, otherwise as an IDENT. Returns `()`. -- seq(*scope₁*, ..., *scopeₙ*): - + parses *scope₁*, ..., *scopeₙ* in this order, and produces a tuple made - out of the parsed values in the same order. As an optimization, all - subscopes of the form STRING are left out of the returned tuple, instead - of returning a useless unit value. It is forbidden for the various - subscopes to refer to the global entry using self of next. - -A few other specific scopes exist to handle Ltac1-like syntax, but their use is -discouraged and they are thus not documented. - -For now there is no way to declare new scopes from Ltac2 side, but this is -planned. - -## Notations - -The Ltac2 parser can be extended by syntactic notations. -``` -VERNAC ::= -| "Ltac2" "Notation" TOKEN+ LEVEL? ":=" TERM - -LEVEL := ":" INT - -TOKEN := -| VAR "(" SCOPE ")" -| STRING -``` - -A Ltac2 notation adds a parsing rule to the Ltac2 grammar, which is expanded -to the provided body where every token from the notation is let-bound to the -corresponding generated expression. - -For instance, assume we perform: -``` -Ltac2 Notation "foo" c(thunk(constr)) ids(list0(ident)) := Bar.f c ids. -``` -Then the following expression -``` -let y := @X in foo (nat -> nat) x $y -``` -will expand at parsing time to -``` -let y := @X in -let c := fun () => constr:(nat -> nat) with ids := [@x; y] in Bar.f c ids -``` - -Beware that the order of evaluation of multiple let-bindings is not specified, -so that you may have to resort to thunking to ensure that side-effects are -performed at the right time. - -## Abbreviations - -There exists a special kind of notations, called abbreviations, that is designed -so that it does not add any parsing rules. It is similar in spirit to Coq -abbreviations, insofar as its main purpose is to give an absolute name to a -piece of pure syntax, which can be transparently referred by this name as if it -were a proper definition. Abbreviations are introduced by the following -syntax. - -``` -VERNAC ::= -| "Ltac2" "Notation" IDENT ":=" TERM -``` - -The abbreviation can then be manipulated just as a normal Ltac2 definition, -except that it is expanded at internalization time into the given expression. -Furthermore, in order to make this kind of construction useful in practice in -an effectful language such as Ltac2, any syntactic argument to an abbreviation -is thunked on-the-fly during its expansion. - -For instance, suppose that we define the following. -``` -Ltac2 Notation foo := fun x => x (). -``` -Then we have the following expansion at internalization time. -``` -foo 0 ↦ (fun x => x ()) (fun _ => 0) -``` - -Note that abbreviations are not typechecked at all, and may result in typing -errors after expansion. - -# Evaluation - -Ltac2 features a toplevel loop that can be used to evaluate expressions. - -``` -VERNAC ::= -| "Ltac2" "Eval" TERM -``` - -This command evaluates the term in the current proof if there is one, or in the -global environment otherwise, and displays the resulting value to the user -together with its type. This function is pure in the sense that it does not -modify the state of the proof, and in particular all side-effects are discarded. - -# Debug - -When the option `Ltac2 Backtrace` is set, toplevel failures will be printed with -a backtrace. - -# Compatibility layer with Ltac1 - -## Ltac1 from Ltac2 - -### Simple API - -One can call Ltac1 code from Ltac2 by using the `ltac1` quotation. It parses -a Ltac1 expression, and semantics of this quotation is the evaluation of the -corresponding code for its side effects. In particular, in cannot return values, -and the quotation has type `unit`. - -Beware, Ltac1 **cannot** access variables from the Ltac2 scope. One is limited -to the use of standalone function calls. - -### Low-level API - -There exists a lower-level FFI into Ltac1 that is not recommended for daily use, -which is available in the `Ltac2.Ltac1` module. This API allows to directly -manipulate dynamically-typed Ltac1 values, either through the function calls, -or using the `ltac1val` quotation. The latter parses the same as `ltac1`, but -has type `Ltac2.Ltac1.t` instead of `unit`, and dynamically behaves as an Ltac1 -thunk, i.e. `ltac1val:(foo)` corresponds to the tactic closure that Ltac1 -would generate from `idtac; foo`. - -Due to intricate dynamic semantics, understanding when Ltac1 value quotations -focus is very hard. This is why some functions return a continuation-passing -style value, as it can dispatch dynamically between focused and unfocused -behaviour. - -## Ltac2 from Ltac1 - -Same as above by switching Ltac1 by Ltac2 and using the `ltac2` quotation -instead. - -Note that the tactic expression is evaluated eagerly, if one wants to use it as -an argument to a Ltac1 function, she has to resort to the good old -`idtac; ltac2:(foo)` trick. For instance, the code below will fail immediately -and won't print anything. - -``` -Ltac mytac tac := idtac "wow"; tac. - -Goal True. -Proof. -mytac ltac2:(fail). -``` - -# Transition from Ltac1 - -Owing to the use of a bunch of notations, the transition shouldn't be -atrociously horrible and shockingly painful up to the point you want to retire -in the Ariège mountains, living off the land and insulting careless bypassers in -proto-georgian. - -That said, we do *not* guarantee you it is going to be a blissful walk either. -Hopefully, owing to the fact Ltac2 is typed, the interactive dialogue with Coq -will help you. - -We list the major changes and the transition strategies hereafter. - -## Syntax changes - -Due to conflicts, a few syntactic rules have changed. - -- The dispatch tactical `tac; [foo|bar]` is now written `tac > [foo|bar]`. -- Levels of a few operators have been revised. Some tacticals now parse as if - they were a normal function, i.e. one has to put parentheses around the - argument when it is complex, e.g an abstraction. List of affected tacticals: - `try`, `repeat`, `do`, `once`, `progress`, `time`, `abstract`. -- `idtac` is no more. Either use `()` if you expect nothing to happen, - `(fun () => ())` if you want a thunk (see next section), or use printing - primitives from the `Message` module if you want to display something. - -## Tactic delay - -Tactics are not magically delayed anymore, neither as functions nor as -arguments. It is your responsibility to thunk them beforehand and apply them -at the call site. - -A typical example of a delayed function: -``` -Ltac foo := blah. -``` -becomes -``` -Ltac2 foo () := blah. -``` - -All subsequent calls to `foo` must be applied to perform the same effect as -before. - -Likewise, for arguments: -``` -Ltac bar tac := tac; tac; tac. -``` -becomes -``` -Ltac2 bar tac := tac (); tac (); tac (). -``` - -We recommend the use of syntactic notations to ease the transition. For -instance, the first example can alternatively written as: -``` -Ltac2 foo0 () := blah. -Ltac2 Notation foo := foo0 (). -``` - -This allows to keep the subsequent calls to the tactic as-is, as the -expression `foo` will be implicitly expanded everywhere into `foo0 ()`. Such -a trick also works for arguments, as arguments of syntactic notations are -implicitly thunked. The second example could thus be written as follows. - -``` -Ltac2 bar0 tac := tac (); tac (); tac (). -Ltac2 Notation bar := bar0. -``` - -## Variable binding - -Ltac1 relies on a crazy amount of dynamic trickery to be able to tell apart -bound variables from terms, hypotheses and whatnot. There is no such thing in -Ltac2, as variables are recognized statically and other constructions do not -live in the same syntactic world. Due to the abuse of quotations, it can -sometimes be complicated to know what a mere identifier represents in a tactic -expression. We recommend tracking the context and letting the compiler spit -typing errors to understand what is going on. - -We list below the typical changes one has to perform depending on the static -errors produced by the typechecker. - -### In Ltac expressions - -- `Unbound value X`, `Unbound constructor X`: - * if `X` is meant to be a term from the current stactic environment, replace - the problematic use by `'X`. - * if `X` is meant to be a hypothesis from the goal context, replace the - problematic use by `&X`. - -### In quotations - -- `The reference X was not found in the current environment`: - * if `X` is meant to be a tactic expression bound by a Ltac2 let or function, - replace the problematic use by `$X`. - * if `X` is meant to be a hypothesis from the goal context, replace the - problematic use by `&X`. - -## Exception catching - -Ltac2 features a proper exception-catching mechanism. For this reason, the -Ltac1 mechanism relying on `fail` taking integers and tacticals decreasing it -has been removed. Now exceptions are preserved by all tacticals, and it is -your duty to catch it and reraise it depending on your use. - -# TODO - -- Implement deep pattern-matching. -- Craft an expressive set of primitive functions -- Implement native compilation to OCaml diff --git a/dune b/dune deleted file mode 100644 index 5dbc4db66a..0000000000 --- a/dune +++ /dev/null @@ -1,3 +0,0 @@ -(env - (dev (flags :standard -rectypes)) - (release (flags :standard -rectypes))) diff --git a/dune-project b/dune-project deleted file mode 100644 index 8154e999de..0000000000 --- a/dune-project +++ /dev/null @@ -1,3 +0,0 @@ -(lang dune 1.6) -(using coq 0.1) -(name ltac2) diff --git a/ltac2.opam b/ltac2.opam deleted file mode 100644 index 47ceb882b1..0000000000 --- a/ltac2.opam +++ /dev/null @@ -1,18 +0,0 @@ -synopsis: "A Tactic Language for Coq." -description: "A Tactic Language for Coq." -name: "coq-ltac2" -opam-version: "2.0" -maintainer: "Pierre-Marie Pédrot " -authors: "Pierre-Marie Pédrot " -homepage: "https://github.com/ppedrot/ltac2" -dev-repo: "https://github.com/ppedrot/ltac2.git" -bug-reports: "https://github.com/ppedrot/ltac2/issues" -license: "LGPL 2.1" -doc: "https://ppedrot.github.io/ltac2/doc" - -depends: [ - "coq" { = "dev" } - "dune" { build & >= "1.9.0" } -] - -build: [ "dune" "build" "-p" name "-j" jobs ] diff --git a/src/dune b/src/dune deleted file mode 100644 index 332f3644b0..0000000000 --- a/src/dune +++ /dev/null @@ -1,11 +0,0 @@ -(library - (name ltac2_plugin) - (public_name ltac2.plugin) - (modules_without_implementation tac2expr tac2qexpr tac2types) - (flags :standard -warn-error -9-27-50) - (libraries coq.plugins.firstorder)) - -(rule - (targets g_ltac2.ml) - (deps (:mlg-file g_ltac2.mlg)) - (action (run coqpp %{mlg-file}))) diff --git a/src/g_ltac2.mlg b/src/g_ltac2.mlg deleted file mode 100644 index 0071dbb088..0000000000 --- a/src/g_ltac2.mlg +++ /dev/null @@ -1,933 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* Tok.t Stream.t -> int option - -let entry_of_lookahead s (lk : lookahead) = - let run strm = match lk 0 strm with None -> err () | Some _ -> () in - Pcoq.Entry.of_parser s run - -let (>>) (lk1 : lookahead) lk2 n strm = match lk1 n strm with -| None -> None -| Some n -> lk2 n strm - -let (<+>) (lk1 : lookahead) lk2 n strm = match lk1 n strm with -| None -> lk2 n strm -| Some n -> Some n - -let lk_kw kw n strm = match stream_nth n strm with -| KEYWORD kw' | IDENT kw' -> if String.equal kw kw' then Some (n + 1) else None -| _ -> None - -let lk_ident n strm = match stream_nth n strm with -| IDENT _ -> Some (n + 1) -| _ -> None - -let lk_int n strm = match stream_nth n strm with -| NUMERAL { NumTok.int = _; frac = ""; exp = "" } -> Some (n + 1) -| _ -> None - -let lk_ident_or_anti = lk_ident <+> (lk_kw "$" >> lk_ident) - -(* lookahead for (x:=t), (?x:=t) and (1:=t) *) -let test_lpar_idnum_coloneq = - entry_of_lookahead "test_lpar_idnum_coloneq" begin - lk_kw "(" >> (lk_ident_or_anti <+> lk_int) >> lk_kw ":=" - end - -(* lookahead for (x:t), (?x:t) *) -let test_lpar_id_colon = - entry_of_lookahead "test_lpar_id_colon" begin - lk_kw "(" >> lk_ident_or_anti >> lk_kw ":" - end - -(* Hack to recognize "(x := t)" and "($x := t)" *) -let test_lpar_id_coloneq = - entry_of_lookahead "test_lpar_id_coloneq" begin - lk_kw "(" >> lk_ident_or_anti >> lk_kw ":=" - end - -(* Hack to recognize "(x)" *) -let test_lpar_id_rpar = - entry_of_lookahead "test_lpar_id_rpar" begin - lk_kw "(" >> lk_ident >> lk_kw ")" - end - -let test_ampersand_ident = - entry_of_lookahead "test_ampersand_ident" begin - lk_kw "&" >> lk_ident - end - -let test_dollar_ident = - entry_of_lookahead "test_dollar_ident" begin - lk_kw "$" >> lk_ident - end - -let tac2expr = Tac2entries.Pltac.tac2expr -let tac2type = Entry.create "tactic:tac2type" -let tac2def_val = Entry.create "tactic:tac2def_val" -let tac2def_typ = Entry.create "tactic:tac2def_typ" -let tac2def_ext = Entry.create "tactic:tac2def_ext" -let tac2def_syn = Entry.create "tactic:tac2def_syn" -let tac2def_mut = Entry.create "tactic:tac2def_mut" -let tac2def_run = Entry.create "tactic:tac2def_run" -let tac2mode = Entry.create "vernac:ltac2_command" - -let ltac1_expr = Pltac.tactic_expr - -let inj_wit wit loc x = CAst.make ~loc @@ CTacExt (wit, x) -let inj_open_constr loc c = inj_wit Tac2quote.wit_open_constr loc c -let inj_pattern loc c = inj_wit Tac2quote.wit_pattern loc c -let inj_reference loc c = inj_wit Tac2quote.wit_reference loc c -let inj_ltac1 loc e = inj_wit Tac2quote.wit_ltac1 loc e -let inj_ltac1val loc e = inj_wit Tac2quote.wit_ltac1val loc e - -let pattern_of_qualid qid = - if Tac2env.is_constructor qid then CAst.make ?loc:qid.CAst.loc @@ CPatRef (RelId qid, []) - else - let open Libnames in - if qualid_is_ident qid then CAst.make ?loc:qid.CAst.loc @@ CPatVar (Name (qualid_basename qid)) - else - CErrors.user_err ?loc:qid.CAst.loc (Pp.str "Syntax error") - -} - -GRAMMAR EXTEND Gram - GLOBAL: tac2expr tac2type tac2def_val tac2def_typ tac2def_ext tac2def_syn - tac2def_mut tac2def_run; - tac2pat: - [ "1" LEFTA - [ qid = Prim.qualid; pl = LIST1 tac2pat LEVEL "0" -> { - if Tac2env.is_constructor qid then - CAst.make ~loc @@ CPatRef (RelId qid, pl) - else - CErrors.user_err ~loc (Pp.str "Syntax error") } - | qid = Prim.qualid -> { pattern_of_qualid qid } - | "["; "]" -> { CAst.make ~loc @@ CPatRef (AbsKn (Other Tac2core.Core.c_nil), []) } - | p1 = tac2pat; "::"; p2 = tac2pat -> - { CAst.make ~loc @@ CPatRef (AbsKn (Other Tac2core.Core.c_cons), [p1; p2])} - ] - | "0" - [ "_" -> { CAst.make ~loc @@ CPatVar Anonymous } - | "()" -> { CAst.make ~loc @@ CPatRef (AbsKn (Tuple 0), []) } - | qid = Prim.qualid -> { pattern_of_qualid qid } - | "("; p = atomic_tac2pat; ")" -> { p } - ] ] - ; - atomic_tac2pat: - [ [ -> - { CAst.make ~loc @@ CPatRef (AbsKn (Tuple 0), []) } - | p = tac2pat; ":"; t = tac2type -> - { CAst.make ~loc @@ CPatCnv (p, t) } - | p = tac2pat; ","; pl = LIST0 tac2pat SEP "," -> - { let pl = p :: pl in - CAst.make ~loc @@ CPatRef (AbsKn (Tuple (List.length pl)), pl) } - | p = tac2pat -> { p } - ] ] - ; - tac2expr: - [ "6" RIGHTA - [ e1 = SELF; ";"; e2 = SELF -> { CAst.make ~loc @@ CTacSeq (e1, e2) } ] - | "5" - [ "fun"; it = LIST1 input_fun ; "=>"; body = tac2expr LEVEL "6" -> - { CAst.make ~loc @@ CTacFun (it, body) } - | "let"; isrec = rec_flag; - lc = LIST1 let_clause SEP "with"; "in"; - e = tac2expr LEVEL "6" -> - { CAst.make ~loc @@ CTacLet (isrec, lc, e) } - | "match"; e = tac2expr LEVEL "5"; "with"; bl = branches; "end" -> - { CAst.make ~loc @@ CTacCse (e, bl) } - ] - | "4" LEFTA [ ] - | "::" RIGHTA - [ e1 = tac2expr; "::"; e2 = tac2expr -> - { CAst.make ~loc @@ CTacApp (CAst.make ~loc @@ CTacCst (AbsKn (Other Tac2core.Core.c_cons)), [e1; e2]) } - ] - | [ e0 = SELF; ","; el = LIST1 NEXT SEP "," -> - { let el = e0 :: el in - CAst.make ~loc @@ CTacApp (CAst.make ~loc @@ CTacCst (AbsKn (Tuple (List.length el))), el) } ] - | "1" LEFTA - [ e = tac2expr; el = LIST1 tac2expr LEVEL "0" -> - { CAst.make ~loc @@ CTacApp (e, el) } - | e = SELF; ".("; qid = Prim.qualid; ")" -> - { CAst.make ~loc @@ CTacPrj (e, RelId qid) } - | e = SELF; ".("; qid = Prim.qualid; ")"; ":="; r = tac2expr LEVEL "5" -> - { CAst.make ~loc @@ CTacSet (e, RelId qid, r) } ] - | "0" - [ "("; a = SELF; ")" -> { a } - | "("; a = SELF; ":"; t = tac2type; ")" -> - { CAst.make ~loc @@ CTacCnv (a, t) } - | "()" -> - { CAst.make ~loc @@ CTacCst (AbsKn (Tuple 0)) } - | "("; ")" -> - { CAst.make ~loc @@ CTacCst (AbsKn (Tuple 0)) } - | "["; a = LIST0 tac2expr LEVEL "5" SEP ";"; "]" -> - { Tac2quote.of_list ~loc (fun x -> x) a } - | "{"; a = tac2rec_fieldexprs; "}" -> - { CAst.make ~loc @@ CTacRec a } - | a = tactic_atom -> { a } ] - ] - ; - branches: - [ [ -> { [] } - | "|"; bl = LIST1 branch SEP "|" -> { bl } - | bl = LIST1 branch SEP "|" -> { bl } ] - ] - ; - branch: - [ [ pat = tac2pat LEVEL "1"; "=>"; e = tac2expr LEVEL "6" -> { (pat, e) } ] ] - ; - rec_flag: - [ [ IDENT "rec" -> { true } - | -> { false } ] ] - ; - mut_flag: - [ [ IDENT "mutable" -> { true } - | -> { false } ] ] - ; - typ_param: - [ [ "'"; id = Prim.ident -> { id } ] ] - ; - tactic_atom: - [ [ n = Prim.integer -> { CAst.make ~loc @@ CTacAtm (AtmInt n) } - | s = Prim.string -> { CAst.make ~loc @@ CTacAtm (AtmStr s) } - | qid = Prim.qualid -> - { if Tac2env.is_constructor qid then - CAst.make ~loc @@ CTacCst (RelId qid) - else - CAst.make ~loc @@ CTacRef (RelId qid) } - | "@"; id = Prim.ident -> { Tac2quote.of_ident (CAst.make ~loc id) } - | "&"; id = lident -> { Tac2quote.of_hyp ~loc id } - | "'"; c = Constr.constr -> { inj_open_constr loc c } - | IDENT "constr"; ":"; "("; c = Constr.lconstr; ")" -> { Tac2quote.of_constr c } - | IDENT "open_constr"; ":"; "("; c = Constr.lconstr; ")" -> { Tac2quote.of_open_constr c } - | IDENT "ident"; ":"; "("; c = lident; ")" -> { Tac2quote.of_ident c } - | IDENT "pattern"; ":"; "("; c = Constr.lconstr_pattern; ")" -> { inj_pattern loc c } - | IDENT "reference"; ":"; "("; c = globref; ")" -> { inj_reference loc c } - | IDENT "ltac1"; ":"; "("; qid = ltac1_expr; ")" -> { inj_ltac1 loc qid } - | IDENT "ltac1val"; ":"; "("; qid = ltac1_expr; ")" -> { inj_ltac1val loc qid } - ] ] - ; - let_clause: - [ [ binder = let_binder; ":="; te = tac2expr -> - { let (pat, fn) = binder in - let te = match fn with - | None -> te - | Some args -> CAst.make ~loc @@ CTacFun (args, te) - in - (pat, te) } - ] ] - ; - let_binder: - [ [ pats = LIST1 input_fun -> - { match pats with - | [{CAst.v=CPatVar _} as pat] -> (pat, None) - | ({CAst.v=CPatVar (Name id)} as pat) :: args -> (pat, Some args) - | [pat] -> (pat, None) - | _ -> CErrors.user_err ~loc (str "Invalid pattern") } - ] ] - ; - tac2type: - [ "5" RIGHTA - [ t1 = tac2type; "->"; t2 = tac2type -> { CAst.make ~loc @@ CTypArrow (t1, t2) } ] - | "2" - [ t = tac2type; "*"; tl = LIST1 tac2type LEVEL "1" SEP "*" -> - { let tl = t :: tl in - CAst.make ~loc @@ CTypRef (AbsKn (Tuple (List.length tl)), tl) } ] - | "1" LEFTA - [ t = SELF; qid = Prim.qualid -> { CAst.make ~loc @@ CTypRef (RelId qid, [t]) } ] - | "0" - [ "("; t = tac2type LEVEL "5"; ")" -> { t } - | id = typ_param -> { CAst.make ~loc @@ CTypVar (Name id) } - | "_" -> { CAst.make ~loc @@ CTypVar Anonymous } - | qid = Prim.qualid -> { CAst.make ~loc @@ CTypRef (RelId qid, []) } - | "("; p = LIST1 tac2type LEVEL "5" SEP ","; ")"; qid = Prim.qualid -> - { CAst.make ~loc @@ CTypRef (RelId qid, p) } ] - ]; - locident: - [ [ id = Prim.ident -> { CAst.make ~loc id } ] ] - ; - binder: - [ [ "_" -> { CAst.make ~loc Anonymous } - | l = Prim.ident -> { CAst.make ~loc (Name l) } ] ] - ; - input_fun: - [ [ b = tac2pat LEVEL "0" -> { b } ] ] - ; - tac2def_body: - [ [ name = binder; it = LIST0 input_fun; ":="; e = tac2expr -> - { let e = if List.is_empty it then e else CAst.make ~loc @@ CTacFun (it, e) in - (name, e) } - ] ] - ; - tac2def_val: - [ [ mut = mut_flag; isrec = rec_flag; l = LIST1 tac2def_body SEP "with" -> - { StrVal (mut, isrec, l) } - ] ] - ; - tac2def_mut: - [ [ "Set"; qid = Prim.qualid; ":="; e = tac2expr -> { StrMut (qid, e) } ] ] - ; - tac2def_run: - [ [ "Eval"; e = tac2expr -> { StrRun e } ] ] - ; - tac2typ_knd: - [ [ t = tac2type -> { CTydDef (Some t) } - | "["; ".."; "]" -> { CTydOpn } - | "["; t = tac2alg_constructors; "]" -> { CTydAlg t } - | "{"; t = tac2rec_fields; "}"-> { CTydRec t } ] ] - ; - tac2alg_constructors: - [ [ "|"; cs = LIST1 tac2alg_constructor SEP "|" -> { cs } - | cs = LIST0 tac2alg_constructor SEP "|" -> { cs } ] ] - ; - tac2alg_constructor: - [ [ c = Prim.ident -> { (c, []) } - | c = Prim.ident; "("; args = LIST0 tac2type SEP ","; ")"-> { (c, args) } ] ] - ; - tac2rec_fields: - [ [ f = tac2rec_field; ";"; l = tac2rec_fields -> { f :: l } - | f = tac2rec_field; ";" -> { [f] } - | f = tac2rec_field -> { [f] } - | -> { [] } ] ] - ; - tac2rec_field: - [ [ mut = mut_flag; id = Prim.ident; ":"; t = tac2type -> { (id, mut, t) } ] ] - ; - tac2rec_fieldexprs: - [ [ f = tac2rec_fieldexpr; ";"; l = tac2rec_fieldexprs -> { f :: l } - | f = tac2rec_fieldexpr; ";" -> { [f] } - | f = tac2rec_fieldexpr-> { [f] } - | -> { [] } ] ] - ; - tac2rec_fieldexpr: - [ [ qid = Prim.qualid; ":="; e = tac2expr LEVEL "1" -> { RelId qid, e } ] ] - ; - tac2typ_prm: - [ [ -> { [] } - | id = typ_param -> { [CAst.make ~loc id] } - | "("; ids = LIST1 [ id = typ_param -> { CAst.make ~loc id } ] SEP "," ;")" -> { ids } - ] ] - ; - tac2typ_def: - [ [ prm = tac2typ_prm; id = Prim.qualid; b = tac2type_body -> { let (r, e) = b in (id, r, (prm, e)) } ] ] - ; - tac2type_body: - [ [ -> { false, CTydDef None } - | ":="; e = tac2typ_knd -> { false, e } - | "::="; e = tac2typ_knd -> { true, e } - ] ] - ; - tac2def_typ: - [ [ "Type"; isrec = rec_flag; l = LIST1 tac2typ_def SEP "with" -> - { StrTyp (isrec, l) } - ] ] - ; - tac2def_ext: - [ [ "@"; IDENT "external"; id = locident; ":"; t = tac2type LEVEL "5"; ":="; - plugin = Prim.string; name = Prim.string -> - { let ml = { mltac_plugin = plugin; mltac_tactic = name } in - StrPrm (id, t, ml) } - ] ] - ; - syn_node: - [ [ "_" -> { CAst.make ~loc None } - | id = Prim.ident -> { CAst.make ~loc (Some id) } - ] ] - ; - sexpr: - [ [ s = Prim.string -> { SexprStr (CAst.make ~loc s) } - | n = Prim.integer -> { SexprInt (CAst.make ~loc n) } - | id = syn_node -> { SexprRec (loc, id, []) } - | id = syn_node; "("; tok = LIST1 sexpr SEP "," ; ")" -> - { SexprRec (loc, id, tok) } - ] ] - ; - syn_level: - [ [ -> { None } - | ":"; n = Prim.integer -> { Some n } - ] ] - ; - tac2def_syn: - [ [ "Notation"; toks = LIST1 sexpr; n = syn_level; ":="; - e = tac2expr -> - { StrSyn (toks, n, e) } - ] ] - ; - lident: - [ [ id = Prim.ident -> { CAst.make ~loc id } ] ] - ; - globref: - [ [ "&"; id = Prim.ident -> { CAst.make ~loc (QHypothesis id) } - | qid = Prim.qualid -> { CAst.make ~loc @@ QReference qid } - ] ] - ; -END - -(* Quotation scopes used by notations *) - -{ - -open Tac2entries.Pltac - -let loc_of_ne_list l = Loc.merge_opt (fst (List.hd l)) (fst (List.last l)) - -} - -GRAMMAR EXTEND Gram - GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause - q_conversion q_rewriting q_clause q_dispatch q_occurrences q_strategy_flag - q_destruction_arg q_reference q_with_bindings q_constr_matching - q_goal_matching q_hintdb q_move_location q_pose q_assert; - anti: - [ [ "$"; id = Prim.ident -> { QAnti (CAst.make ~loc id) } ] ] - ; - ident_or_anti: - [ [ id = lident -> { QExpr id } - | "$"; id = Prim.ident -> { QAnti (CAst.make ~loc id) } - ] ] - ; - lident: - [ [ id = Prim.ident -> { CAst.make ~loc id } ] ] - ; - lnatural: - [ [ n = Prim.natural -> { CAst.make ~loc n } ] ] - ; - q_ident: - [ [ id = ident_or_anti -> { id } ] ] - ; - qhyp: - [ [ x = anti -> { x } - | n = lnatural -> { QExpr (CAst.make ~loc @@ QAnonHyp n) } - | id = lident -> { QExpr (CAst.make ~loc @@ QNamedHyp id) } - ] ] - ; - simple_binding: - [ [ "("; h = qhyp; ":="; c = Constr.lconstr; ")" -> - { CAst.make ~loc (h, c) } - ] ] - ; - bindings: - [ [ test_lpar_idnum_coloneq; bl = LIST1 simple_binding -> - { CAst.make ~loc @@ QExplicitBindings bl } - | bl = LIST1 Constr.constr -> - { CAst.make ~loc @@ QImplicitBindings bl } - ] ] - ; - q_bindings: - [ [ bl = bindings -> { bl } ] ] - ; - q_with_bindings: - [ [ bl = with_bindings -> { bl } ] ] - ; - intropatterns: - [ [ l = LIST0 nonsimple_intropattern -> { CAst.make ~loc l } ] ] - ; -(* ne_intropatterns: *) -(* [ [ l = LIST1 nonsimple_intropattern -> l ]] *) -(* ; *) - or_and_intropattern: - [ [ "["; tc = LIST1 intropatterns SEP "|"; "]" -> { CAst.make ~loc @@ QIntroOrPattern tc } - | "()" -> { CAst.make ~loc @@ QIntroAndPattern (CAst.make ~loc []) } - | "("; si = simple_intropattern; ")" -> { CAst.make ~loc @@ QIntroAndPattern (CAst.make ~loc [si]) } - | "("; si = simple_intropattern; ","; - tc = LIST1 simple_intropattern SEP "," ; ")" -> - { CAst.make ~loc @@ QIntroAndPattern (CAst.make ~loc (si::tc)) } - | "("; si = simple_intropattern; "&"; - tc = LIST1 simple_intropattern SEP "&" ; ")" -> - (* (A & B & C) is translated into (A,(B,C)) *) - { let rec pairify = function - | ([]|[_]|[_;_]) as l -> CAst.make ~loc l - | t::q -> - let q = - CAst.make ~loc @@ - QIntroAction (CAst.make ~loc @@ - QIntroOrAndPattern (CAst.make ~loc @@ - QIntroAndPattern (pairify q))) - in - CAst.make ~loc [t; q] - in CAst.make ~loc @@ QIntroAndPattern (pairify (si::tc)) } ] ] - ; - equality_intropattern: - [ [ "->" -> { CAst.make ~loc @@ QIntroRewrite true } - | "<-" -> { CAst.make ~loc @@ QIntroRewrite false } - | "[="; tc = intropatterns; "]" -> { CAst.make ~loc @@ QIntroInjection tc } ] ] - ; - naming_intropattern: - [ [ LEFTQMARK; id = lident -> - { CAst.make ~loc @@ QIntroFresh (QExpr id) } - | "?$"; id = lident -> - { CAst.make ~loc @@ QIntroFresh (QAnti id) } - | "?" -> - { CAst.make ~loc @@ QIntroAnonymous } - | id = ident_or_anti -> - { CAst.make ~loc @@ QIntroIdentifier id } - ] ] - ; - nonsimple_intropattern: - [ [ l = simple_intropattern -> { l } - | "*" -> { CAst.make ~loc @@ QIntroForthcoming true } - | "**" -> { CAst.make ~loc @@ QIntroForthcoming false } ] ] - ; - simple_intropattern: - [ [ pat = simple_intropattern_closed -> -(* l = LIST0 ["%"; c = operconstr LEVEL "0" -> c] -> *) - (** TODO: handle %pat *) - { pat } - ] ] - ; - simple_intropattern_closed: - [ [ pat = or_and_intropattern -> - { CAst.make ~loc @@ QIntroAction (CAst.make ~loc @@ QIntroOrAndPattern pat) } - | pat = equality_intropattern -> - { CAst.make ~loc @@ QIntroAction pat } - | "_" -> - { CAst.make ~loc @@ QIntroAction (CAst.make ~loc @@ QIntroWildcard) } - | pat = naming_intropattern -> - { CAst.make ~loc @@ QIntroNaming pat } - ] ] - ; - q_intropatterns: - [ [ ipat = intropatterns -> { ipat } ] ] - ; - q_intropattern: - [ [ ipat = simple_intropattern -> { ipat } ] ] - ; - nat_or_anti: - [ [ n = lnatural -> { QExpr n } - | "$"; id = Prim.ident -> { QAnti (CAst.make ~loc id) } - ] ] - ; - eqn_ipat: - [ [ IDENT "eqn"; ":"; pat = naming_intropattern -> { Some pat } - | -> { None } - ] ] - ; - with_bindings: - [ [ "with"; bl = bindings -> { bl } | -> { CAst.make ~loc @@ QNoBindings } ] ] - ; - constr_with_bindings: - [ [ c = Constr.constr; l = with_bindings -> { CAst.make ~loc @@ (c, l) } ] ] - ; - destruction_arg: - [ [ n = lnatural -> { CAst.make ~loc @@ QElimOnAnonHyp n } - | id = lident -> { CAst.make ~loc @@ QElimOnIdent id } - | c = constr_with_bindings -> { CAst.make ~loc @@ QElimOnConstr c } - ] ] - ; - q_destruction_arg: - [ [ arg = destruction_arg -> { arg } ] ] - ; - as_or_and_ipat: - [ [ "as"; ipat = or_and_intropattern -> { Some ipat } - | -> { None } - ] ] - ; - occs_nums: - [ [ nl = LIST1 nat_or_anti -> { CAst.make ~loc @@ QOnlyOccurrences nl } - | "-"; n = nat_or_anti; nl = LIST0 nat_or_anti -> - { CAst.make ~loc @@ QAllOccurrencesBut (n::nl) } - ] ] - ; - occs: - [ [ "at"; occs = occs_nums -> { occs } | -> { CAst.make ~loc QAllOccurrences } ] ] - ; - hypident: - [ [ id = ident_or_anti -> - { id,Locus.InHyp } - | "("; IDENT "type"; IDENT "of"; id = ident_or_anti; ")" -> - { id,Locus.InHypTypeOnly } - | "("; IDENT "value"; IDENT "of"; id = ident_or_anti; ")" -> - { id,Locus.InHypValueOnly } - ] ] - ; - hypident_occ: - [ [ h=hypident; occs=occs -> { let (id,l) = h in ((occs,id),l) } ] ] - ; - in_clause: - [ [ "*"; occs=occs -> - { { q_onhyps = None; q_concl_occs = occs } } - | "*"; "|-"; occs = concl_occ -> - { { q_onhyps = None; q_concl_occs = occs } } - | hl = LIST0 hypident_occ SEP ","; "|-"; occs = concl_occ -> - { { q_onhyps = Some hl; q_concl_occs = occs } } - | hl = LIST0 hypident_occ SEP "," -> - { { q_onhyps = Some hl; q_concl_occs = CAst.make ~loc QNoOccurrences } } - ] ] - ; - clause: - [ [ "in"; cl = in_clause -> { CAst.make ~loc @@ cl } - | "at"; occs = occs_nums -> - { CAst.make ~loc @@ { q_onhyps = Some []; q_concl_occs = occs } } - ] ] - ; - q_clause: - [ [ cl = clause -> { cl } ] ] - ; - concl_occ: - [ [ "*"; occs = occs -> { occs } - | -> { CAst.make ~loc QNoOccurrences } - ] ] - ; - induction_clause: - [ [ c = destruction_arg; pat = as_or_and_ipat; eq = eqn_ipat; - cl = OPT clause -> - { CAst.make ~loc @@ { - indcl_arg = c; - indcl_eqn = eq; - indcl_as = pat; - indcl_in = cl; - } } - ] ] - ; - q_induction_clause: - [ [ cl = induction_clause -> { cl } ] ] - ; - conversion: - [ [ c = Constr.constr -> - { CAst.make ~loc @@ QConvert c } - | c1 = Constr.constr; "with"; c2 = Constr.constr -> - { CAst.make ~loc @@ QConvertWith (c1, c2) } - ] ] - ; - q_conversion: - [ [ c = conversion -> { c } ] ] - ; - orient: - [ [ "->" -> { CAst.make ~loc (Some true) } - | "<-" -> { CAst.make ~loc (Some false) } - | -> { CAst.make ~loc None } - ]] - ; - rewriter: - [ [ "!"; c = constr_with_bindings -> - { (CAst.make ~loc @@ QRepeatPlus,c) } - | [ "?" -> { () } | LEFTQMARK -> { () } ]; c = constr_with_bindings -> - { (CAst.make ~loc @@ QRepeatStar,c) } - | n = lnatural; "!"; c = constr_with_bindings -> - { (CAst.make ~loc @@ QPrecisely n,c) } - | n = lnatural; ["?" -> { () } | LEFTQMARK -> { () } ]; c = constr_with_bindings -> - { (CAst.make ~loc @@ QUpTo n,c) } - | n = lnatural; c = constr_with_bindings -> - { (CAst.make ~loc @@ QPrecisely n,c) } - | c = constr_with_bindings -> - { (CAst.make ~loc @@ QPrecisely (CAst.make 1), c) } - ] ] - ; - oriented_rewriter: - [ [ b = orient; r = rewriter -> - { let (m, c) = r in - CAst.make ~loc @@ { - rew_orient = b; - rew_repeat = m; - rew_equatn = c; - } } - ] ] - ; - q_rewriting: - [ [ r = oriented_rewriter -> { r } ] ] - ; - tactic_then_last: - [ [ "|"; lta = LIST0 (OPT tac2expr LEVEL "6") SEP "|" -> { lta } - | -> { [] } - ] ] - ; - tactic_then_gen: - [ [ ta = tac2expr; "|"; tg = tactic_then_gen -> { let (first,last) = tg in (Some ta :: first, last) } - | ta = tac2expr; ".."; l = tactic_then_last -> { ([], Some (Some ta, l)) } - | ".."; l = tactic_then_last -> { ([], Some (None, l)) } - | ta = tac2expr -> { ([Some ta], None) } - | "|"; tg = tactic_then_gen -> { let (first,last) = tg in (None :: first, last) } - | -> { ([None], None) } - ] ] - ; - q_dispatch: - [ [ d = tactic_then_gen -> { CAst.make ~loc d } ] ] - ; - q_occurrences: - [ [ occs = occs -> { occs } ] ] - ; - red_flag: - [ [ IDENT "beta" -> { CAst.make ~loc @@ QBeta } - | IDENT "iota" -> { CAst.make ~loc @@ QIota } - | IDENT "match" -> { CAst.make ~loc @@ QMatch } - | IDENT "fix" -> { CAst.make ~loc @@ QFix } - | IDENT "cofix" -> { CAst.make ~loc @@ QCofix } - | IDENT "zeta" -> { CAst.make ~loc @@ QZeta } - | IDENT "delta"; d = delta_flag -> { d } - ] ] - ; - refglobal: - [ [ "&"; id = Prim.ident -> { QExpr (CAst.make ~loc @@ QHypothesis id) } - | qid = Prim.qualid -> { QExpr (CAst.make ~loc @@ QReference qid) } - | "$"; id = Prim.ident -> { QAnti (CAst.make ~loc id) } - ] ] - ; - q_reference: - [ [ r = refglobal -> { r } ] ] - ; - refglobals: - [ [ gl = LIST1 refglobal -> { CAst.make ~loc gl } ] ] - ; - delta_flag: - [ [ "-"; "["; idl = refglobals; "]" -> { CAst.make ~loc @@ QDeltaBut idl } - | "["; idl = refglobals; "]" -> { CAst.make ~loc @@ QConst idl } - | -> { CAst.make ~loc @@ QDeltaBut (CAst.make ~loc []) } - ] ] - ; - strategy_flag: - [ [ s = LIST1 red_flag -> { CAst.make ~loc s } - | d = delta_flag -> - { CAst.make ~loc - [CAst.make ~loc QBeta; CAst.make ~loc QIota; CAst.make ~loc QZeta; d] } - ] ] - ; - q_strategy_flag: - [ [ flag = strategy_flag -> { flag } ] ] - ; - hintdb: - [ [ "*" -> { CAst.make ~loc @@ QHintAll } - | l = LIST1 ident_or_anti -> { CAst.make ~loc @@ QHintDbs l } - ] ] - ; - q_hintdb: - [ [ db = hintdb -> { db } ] ] - ; - match_pattern: - [ [ IDENT "context"; id = OPT Prim.ident; - "["; pat = Constr.lconstr_pattern; "]" -> { CAst.make ~loc @@ QConstrMatchContext (id, pat) } - | pat = Constr.lconstr_pattern -> { CAst.make ~loc @@ QConstrMatchPattern pat } ] ] - ; - match_rule: - [ [ mp = match_pattern; "=>"; tac = tac2expr -> - { CAst.make ~loc @@ (mp, tac) } - ] ] - ; - match_list: - [ [ mrl = LIST1 match_rule SEP "|" -> { CAst.make ~loc @@ mrl } - | "|"; mrl = LIST1 match_rule SEP "|" -> { CAst.make ~loc @@ mrl } ] ] - ; - q_constr_matching: - [ [ m = match_list -> { m } ] ] - ; - gmatch_hyp_pattern: - [ [ na = Prim.name; ":"; pat = match_pattern -> { (na, pat) } ] ] - ; - gmatch_pattern: - [ [ "["; hl = LIST0 gmatch_hyp_pattern SEP ","; "|-"; p = match_pattern; "]" -> - { CAst.make ~loc @@ { - q_goal_match_concl = p; - q_goal_match_hyps = hl; - } } - ] ] - ; - gmatch_rule: - [ [ mp = gmatch_pattern; "=>"; tac = tac2expr -> - { CAst.make ~loc @@ (mp, tac) } - ] ] - ; - gmatch_list: - [ [ mrl = LIST1 gmatch_rule SEP "|" -> { CAst.make ~loc @@ mrl } - | "|"; mrl = LIST1 gmatch_rule SEP "|" -> { CAst.make ~loc @@ mrl } ] ] - ; - q_goal_matching: - [ [ m = gmatch_list -> { m } ] ] - ; - move_location: - [ [ "at"; IDENT "top" -> { CAst.make ~loc @@ QMoveFirst } - | "at"; IDENT "bottom" -> { CAst.make ~loc @@ QMoveLast } - | IDENT "after"; id = ident_or_anti -> { CAst.make ~loc @@ QMoveAfter id } - | IDENT "before"; id = ident_or_anti -> { CAst.make ~loc @@ QMoveBefore id } - ] ] - ; - q_move_location: - [ [ mv = move_location -> { mv } ] ] - ; - as_name: - [ [ -> { None } - | "as"; id = ident_or_anti -> { Some id } - ] ] - ; - pose: - [ [ test_lpar_id_coloneq; "("; id = ident_or_anti; ":="; c = Constr.lconstr; ")" -> - { CAst.make ~loc (Some id, c) } - | c = Constr.constr; na = as_name -> { CAst.make ~loc (na, c) } - ] ] - ; - q_pose: - [ [ p = pose -> { p } ] ] - ; - as_ipat: - [ [ "as"; ipat = simple_intropattern -> { Some ipat } - | -> { None } - ] ] - ; - by_tactic: - [ [ "by"; tac = tac2expr -> { Some tac } - | -> { None } - ] ] - ; - assertion: - [ [ test_lpar_id_coloneq; "("; id = ident_or_anti; ":="; c = Constr.lconstr; ")" -> - { CAst.make ~loc (QAssertValue (id, c)) } - | test_lpar_id_colon; "("; id = ident_or_anti; ":"; c = Constr.lconstr; ")"; tac = by_tactic -> - { let ipat = CAst.make ~loc @@ QIntroNaming (CAst.make ~loc @@ QIntroIdentifier id) in - CAst.make ~loc (QAssertType (Some ipat, c, tac)) } - | c = Constr.constr; ipat = as_ipat; tac = by_tactic -> - { CAst.make ~loc (QAssertType (ipat, c, tac)) } - ] ] - ; - q_assert: - [ [ a = assertion -> { a } ] ] - ; -END - -(** Extension of constr syntax *) - -(* -GRAMMAR EXTEND Gram - Pcoq.Constr.operconstr: LEVEL "0" - [ [ IDENT "ltac2"; ":"; "("; tac = tac2expr; ")" -> - { let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in - CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) } - | test_ampersand_ident; "&"; id = Prim.ident -> - { let tac = Tac2quote.of_exact_hyp ~loc (CAst.make ~loc id) in - let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in - CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) } - | test_dollar_ident; "$"; id = Prim.ident -> - { let id = Loc.tag ~loc id in - let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2_quotation) id in - CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) } - ] ] - ; -END -*) -{ - -let () = - -let open Extend in -let open Tok in -let (++) r s = Next (r, s) in -let rules = [ - Rule ( - Stop ++ Aentry test_dollar_ident ++ Atoken (PKEYWORD "$") ++ Aentry Prim.ident, - begin fun id _ _ loc -> - let id = Loc.tag ~loc id in - let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2_quotation) id in - CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) - end - ); - - Rule ( - Stop ++ Aentry test_ampersand_ident ++ Atoken (PKEYWORD "&") ++ Aentry Prim.ident, - begin fun id _ _ loc -> - let tac = Tac2quote.of_exact_hyp ~loc (CAst.make ~loc id) in - let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in - CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) - end - ); - - Rule ( - Stop ++ Atoken (PIDENT (Some "ltac2")) ++ Atoken (PKEYWORD ":") ++ - Atoken (PKEYWORD "(") ++ Aentry tac2expr ++ Atoken (PKEYWORD ")"), - begin fun _ tac _ _ _ loc -> - let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in - CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) - end - ) -] in - -Hook.set Tac2entries.register_constr_quotations begin fun () -> - Pcoq.grammar_extend Pcoq.Constr.operconstr None (Some (Gramlib.Gramext.Level "0"), [(None, None, rules)]) -end - -} - -{ - -let pr_ltac2entry _ = mt () (* FIXME *) -let pr_ltac2expr _ = mt () (* FIXME *) - -} - -VERNAC ARGUMENT EXTEND ltac2_entry -PRINTED BY { pr_ltac2entry } -| [ tac2def_val(v) ] -> { v } -| [ tac2def_typ(t) ] -> { t } -| [ tac2def_ext(e) ] -> { e } -| [ tac2def_syn(e) ] -> { e } -| [ tac2def_mut(e) ] -> { e } -| [ tac2def_run(e) ] -> { e } -END - -{ - -let classify_ltac2 = function -| StrSyn _ -> Vernacextend.(VtSideff [], VtNow) -| StrMut _ | StrVal _ | StrPrm _ | StrTyp _ | StrRun _ -> Vernacextend.classify_as_sideeff - -} - -VERNAC COMMAND EXTEND VernacDeclareTactic2Definition -| #[ local = locality ] ![proof] [ "Ltac2" ltac2_entry(e) ] => { classify_ltac2 e } -> { - fun ~pstate -> Tac2entries.register_struct ?local ~pstate e; pstate - } -END - -{ - -let _ = Pvernac.register_proof_mode "Ltac2" tac2mode - -} - -VERNAC ARGUMENT EXTEND ltac2_expr -PRINTED BY { pr_ltac2expr } -| [ tac2expr(e) ] -> { e } -END - -{ - -open G_ltac -open Vernacextend - -} - -VERNAC { tac2mode } EXTEND VernacLtac2 -| ![proof] [ ltac2_expr(t) ltac_use_default(default) ] => - { classify_as_proofstep } -> { -(* let g = Option.default (Proof_global.get_default_goal_selector ()) g in *) - fun ~pstate -> - Option.map (fun pstate -> Tac2entries.call ~pstate ~default t) pstate - } -END - -{ - -open Stdarg - -} - -VERNAC COMMAND EXTEND Ltac2Print CLASSIFIED AS SIDEFF -| [ "Print" "Ltac2" reference(tac) ] -> { Tac2entries.print_ltac tac } -END diff --git a/src/ltac2_plugin.mlpack b/src/ltac2_plugin.mlpack deleted file mode 100644 index 2a25e825cb..0000000000 --- a/src/ltac2_plugin.mlpack +++ /dev/null @@ -1,14 +0,0 @@ -Tac2dyn -Tac2ffi -Tac2env -Tac2print -Tac2intern -Tac2interp -Tac2entries -Tac2quote -Tac2match -Tac2core -Tac2extffi -Tac2tactics -Tac2stdlib -G_ltac2 diff --git a/src/tac2core.ml b/src/tac2core.ml deleted file mode 100644 index d7e7b91ee6..0000000000 --- a/src/tac2core.ml +++ /dev/null @@ -1,1446 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* Value.of_option Value.of_ident None -| Name id -> Value.of_option Value.of_ident (Some id) - -let to_name c = match Value.to_option Value.to_ident c with -| None -> Anonymous -| Some id -> Name id - -let of_relevance = function - | Sorts.Relevant -> ValInt 0 - | Sorts.Irrelevant -> ValInt 1 - -let to_relevance = function - | ValInt 0 -> Sorts.Relevant - | ValInt 1 -> Sorts.Irrelevant - | _ -> assert false - -let of_annot f Context.{binder_name;binder_relevance} = - of_tuple [|(f binder_name); of_relevance binder_relevance|] - -let to_annot f x = - match to_tuple x with - | [|x;y|] -> - let x = f x in - let y = to_relevance y in - Context.make_annot x y - | _ -> assert false - -let of_instance u = - let u = Univ.Instance.to_array (EConstr.Unsafe.to_instance u) in - Value.of_array (fun v -> Value.of_ext Value.val_univ v) u - -let to_instance u = - let u = Value.to_array (fun v -> Value.to_ext Value.val_univ v) u in - EConstr.EInstance.make (Univ.Instance.of_array u) - -let of_rec_declaration (nas, ts, cs) = - (Value.of_array (of_annot of_name) nas, - Value.of_array Value.of_constr ts, - Value.of_array Value.of_constr cs) - -let to_rec_declaration (nas, ts, cs) = - (Value.to_array (to_annot to_name) nas, - Value.to_array Value.to_constr ts, - Value.to_array Value.to_constr cs) - -let of_result f = function -| Inl c -> v_blk 0 [|f c|] -| Inr e -> v_blk 1 [|Value.of_exn e|] - -(** Stdlib exceptions *) - -let err_notfocussed = - Tac2interp.LtacError (coq_core "Not_focussed", [||]) - -let err_outofbounds = - Tac2interp.LtacError (coq_core "Out_of_bounds", [||]) - -let err_notfound = - Tac2interp.LtacError (coq_core "Not_found", [||]) - -let err_matchfailure = - Tac2interp.LtacError (coq_core "Match_failure", [||]) - -(** Helper functions *) - -let thaw f = Tac2ffi.apply f [v_unit] - -let fatal_flag : unit Exninfo.t = Exninfo.make () - -let set_bt info = - if !Tac2interp.print_ltac2_backtrace then - Tac2interp.get_backtrace >>= fun bt -> - Proofview.tclUNIT (Exninfo.add info Tac2entries.backtrace bt) - else Proofview.tclUNIT info - -let throw ?(info = Exninfo.null) e = - set_bt info >>= fun info -> - let info = Exninfo.add info fatal_flag () in - Proofview.tclLIFT (Proofview.NonLogical.raise ~info e) - -let fail ?(info = Exninfo.null) e = - set_bt info >>= fun info -> - Proofview.tclZERO ~info e - -let return x = Proofview.tclUNIT x -let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } - -let wrap f = - return () >>= fun () -> return (f ()) - -let wrap_unit f = - return () >>= fun () -> f (); return v_unit - -let assert_focussed = - Proofview.Goal.goals >>= fun gls -> - match gls with - | [_] -> Proofview.tclUNIT () - | [] | _ :: _ :: _ -> throw err_notfocussed - -let pf_apply f = - Proofview.Goal.goals >>= function - | [] -> - Proofview.tclENV >>= fun env -> - Proofview.tclEVARMAP >>= fun sigma -> - f env sigma - | [gl] -> - gl >>= fun gl -> - f (Proofview.Goal.env gl) (Tacmach.New.project gl) - | _ :: _ :: _ -> - throw err_notfocussed - -(** Primitives *) - -let define_primitive name arity f = - Tac2env.define_primitive (pname name) (mk_closure arity f) - -let define0 name f = define_primitive name arity_one (fun _ -> f) - -let define1 name r0 f = define_primitive name arity_one begin fun x -> - f (Value.repr_to r0 x) -end - -let define2 name r0 r1 f = define_primitive name (arity_suc arity_one) begin fun x y -> - f (Value.repr_to r0 x) (Value.repr_to r1 y) -end - -let define3 name r0 r1 r2 f = define_primitive name (arity_suc (arity_suc arity_one)) begin fun x y z -> - f (Value.repr_to r0 x) (Value.repr_to r1 y) (Value.repr_to r2 z) -end - -(** Printing *) - -let () = define1 "print" pp begin fun pp -> - wrap_unit (fun () -> Feedback.msg_notice pp) -end - -let () = define1 "message_of_int" int begin fun n -> - return (Value.of_pp (Pp.int n)) -end - -let () = define1 "message_of_string" string begin fun s -> - return (Value.of_pp (str (Bytes.to_string s))) -end - -let () = define1 "message_of_constr" constr begin fun c -> - pf_apply begin fun env sigma -> - let pp = Printer.pr_econstr_env env sigma c in - return (Value.of_pp pp) - end -end - -let () = define1 "message_of_ident" ident begin fun c -> - let pp = Id.print c in - return (Value.of_pp pp) -end - -let () = define1 "message_of_exn" valexpr begin fun v -> - Proofview.tclENV >>= fun env -> - Proofview.tclEVARMAP >>= fun sigma -> - let pp = Tac2print.pr_valexpr env sigma v (GTypRef (Other Core.t_exn, [])) in - return (Value.of_pp pp) -end - - -let () = define2 "message_concat" pp pp begin fun m1 m2 -> - return (Value.of_pp (Pp.app m1 m2)) -end - -(** Array *) - -let () = define2 "array_make" int valexpr begin fun n x -> - if n < 0 || n > Sys.max_array_length then throw err_outofbounds - else wrap (fun () -> v_blk 0 (Array.make n x)) -end - -let () = define1 "array_length" block begin fun (_, v) -> - return (Value.of_int (Array.length v)) -end - -let () = define3 "array_set" block int valexpr begin fun (_, v) n x -> - if n < 0 || n >= Array.length v then throw err_outofbounds - else wrap_unit (fun () -> v.(n) <- x) -end - -let () = define2 "array_get" block int begin fun (_, v) n -> - if n < 0 || n >= Array.length v then throw err_outofbounds - else wrap (fun () -> v.(n)) -end - -(** Ident *) - -let () = define2 "ident_equal" ident ident begin fun id1 id2 -> - return (Value.of_bool (Id.equal id1 id2)) -end - -let () = define1 "ident_to_string" ident begin fun id -> - return (Value.of_string (Bytes.of_string (Id.to_string id))) -end - -let () = define1 "ident_of_string" string begin fun s -> - let id = try Some (Id.of_string (Bytes.to_string s)) with _ -> None in - return (Value.of_option Value.of_ident id) -end - -(** Int *) - -let () = define2 "int_equal" int int begin fun m n -> - return (Value.of_bool (m == n)) -end - -let binop n f = define2 n int int begin fun m n -> - return (Value.of_int (f m n)) -end - -let () = binop "int_compare" Int.compare -let () = binop "int_add" (+) -let () = binop "int_sub" (-) -let () = binop "int_mul" ( * ) - -let () = define1 "int_neg" int begin fun m -> - return (Value.of_int (~- m)) -end - -(** Char *) - -let () = define1 "char_of_int" int begin fun n -> - wrap (fun () -> Value.of_char (Char.chr n)) -end - -let () = define1 "char_to_int" char begin fun n -> - wrap (fun () -> Value.of_int (Char.code n)) -end - -(** String *) - -let () = define2 "string_make" int char begin fun n c -> - if n < 0 || n > Sys.max_string_length then throw err_outofbounds - else wrap (fun () -> Value.of_string (Bytes.make n c)) -end - -let () = define1 "string_length" string begin fun s -> - return (Value.of_int (Bytes.length s)) -end - -let () = define3 "string_set" string int char begin fun s n c -> - if n < 0 || n >= Bytes.length s then throw err_outofbounds - else wrap_unit (fun () -> Bytes.set s n c) -end - -let () = define2 "string_get" string int begin fun s n -> - if n < 0 || n >= Bytes.length s then throw err_outofbounds - else wrap (fun () -> Value.of_char (Bytes.get s n)) -end - -(** Terms *) - -(** constr -> constr *) -let () = define1 "constr_type" constr begin fun c -> - let get_type env sigma = - Proofview.V82.wrap_exceptions begin fun () -> - let (sigma, t) = Typing.type_of env sigma c in - let t = Value.of_constr t in - Proofview.Unsafe.tclEVARS sigma <*> Proofview.tclUNIT t - end in - pf_apply get_type -end - -(** constr -> constr *) -let () = define2 "constr_equal" constr constr begin fun c1 c2 -> - Proofview.tclEVARMAP >>= fun sigma -> - let b = EConstr.eq_constr sigma c1 c2 in - Proofview.tclUNIT (Value.of_bool b) -end - -let () = define1 "constr_kind" constr begin fun c -> - let open Constr in - Proofview.tclEVARMAP >>= fun sigma -> - return begin match EConstr.kind sigma c with - | Rel n -> - v_blk 0 [|Value.of_int n|] - | Var id -> - v_blk 1 [|Value.of_ident id|] - | Meta n -> - v_blk 2 [|Value.of_int n|] - | Evar (evk, args) -> - v_blk 3 [| - Value.of_int (Evar.repr evk); - Value.of_array Value.of_constr args; - |] - | Sort s -> - v_blk 4 [|Value.of_ext Value.val_sort s|] - | Cast (c, k, t) -> - v_blk 5 [| - Value.of_constr c; - Value.of_ext Value.val_cast k; - Value.of_constr t; - |] - | Prod (na, t, u) -> - v_blk 6 [| - of_annot of_name na; - Value.of_constr t; - Value.of_constr u; - |] - | Lambda (na, t, c) -> - v_blk 7 [| - of_annot of_name na; - Value.of_constr t; - Value.of_constr c; - |] - | LetIn (na, b, t, c) -> - v_blk 8 [| - of_annot of_name na; - Value.of_constr b; - Value.of_constr t; - Value.of_constr c; - |] - | App (c, cl) -> - v_blk 9 [| - Value.of_constr c; - Value.of_array Value.of_constr cl; - |] - | Const (cst, u) -> - v_blk 10 [| - Value.of_constant cst; - of_instance u; - |] - | Ind (ind, u) -> - v_blk 11 [| - Value.of_ext Value.val_inductive ind; - of_instance u; - |] - | Construct (cstr, u) -> - v_blk 12 [| - Value.of_ext Value.val_constructor cstr; - of_instance u; - |] - | Case (ci, c, t, bl) -> - v_blk 13 [| - Value.of_ext Value.val_case ci; - Value.of_constr c; - Value.of_constr t; - Value.of_array Value.of_constr bl; - |] - | Fix ((recs, i), def) -> - let (nas, ts, cs) = of_rec_declaration def in - v_blk 14 [| - Value.of_array Value.of_int recs; - Value.of_int i; - nas; - ts; - cs; - |] - | CoFix (i, def) -> - let (nas, ts, cs) = of_rec_declaration def in - v_blk 15 [| - Value.of_int i; - nas; - ts; - cs; - |] - | Proj (p, c) -> - v_blk 16 [| - Value.of_ext Value.val_projection p; - Value.of_constr c; - |] - | Int _ -> - assert false - end -end - -let () = define1 "constr_make" valexpr begin fun knd -> - let c = match Tac2ffi.to_block knd with - | (0, [|n|]) -> - let n = Value.to_int n in - EConstr.mkRel n - | (1, [|id|]) -> - let id = Value.to_ident id in - EConstr.mkVar id - | (2, [|n|]) -> - let n = Value.to_int n in - EConstr.mkMeta n - | (3, [|evk; args|]) -> - let evk = Evar.unsafe_of_int (Value.to_int evk) in - let args = Value.to_array Value.to_constr args in - EConstr.mkEvar (evk, args) - | (4, [|s|]) -> - let s = Value.to_ext Value.val_sort s in - EConstr.mkSort (EConstr.Unsafe.to_sorts s) - | (5, [|c; k; t|]) -> - let c = Value.to_constr c in - let k = Value.to_ext Value.val_cast k in - let t = Value.to_constr t in - EConstr.mkCast (c, k, t) - | (6, [|na; t; u|]) -> - let na = to_annot to_name na in - let t = Value.to_constr t in - let u = Value.to_constr u in - EConstr.mkProd (na, t, u) - | (7, [|na; t; c|]) -> - let na = to_annot to_name na in - let t = Value.to_constr t in - let u = Value.to_constr c in - EConstr.mkLambda (na, t, u) - | (8, [|na; b; t; c|]) -> - let na = to_annot to_name na in - let b = Value.to_constr b in - let t = Value.to_constr t in - let c = Value.to_constr c in - EConstr.mkLetIn (na, b, t, c) - | (9, [|c; cl|]) -> - let c = Value.to_constr c in - let cl = Value.to_array Value.to_constr cl in - EConstr.mkApp (c, cl) - | (10, [|cst; u|]) -> - let cst = Value.to_constant cst in - let u = to_instance u in - EConstr.mkConstU (cst, u) - | (11, [|ind; u|]) -> - let ind = Value.to_ext Value.val_inductive ind in - let u = to_instance u in - EConstr.mkIndU (ind, u) - | (12, [|cstr; u|]) -> - let cstr = Value.to_ext Value.val_constructor cstr in - let u = to_instance u in - EConstr.mkConstructU (cstr, u) - | (13, [|ci; c; t; bl|]) -> - let ci = Value.to_ext Value.val_case ci in - let c = Value.to_constr c in - let t = Value.to_constr t in - let bl = Value.to_array Value.to_constr bl in - EConstr.mkCase (ci, c, t, bl) - | (14, [|recs; i; nas; ts; cs|]) -> - let recs = Value.to_array Value.to_int recs in - let i = Value.to_int i in - let def = to_rec_declaration (nas, ts, cs) in - EConstr.mkFix ((recs, i), def) - | (15, [|i; nas; ts; cs|]) -> - let i = Value.to_int i in - let def = to_rec_declaration (nas, ts, cs) in - EConstr.mkCoFix (i, def) - | (16, [|p; c|]) -> - let p = Value.to_ext Value.val_projection p in - let c = Value.to_constr c in - EConstr.mkProj (p, c) - | _ -> assert false - in - return (Value.of_constr c) -end - -let () = define1 "constr_check" constr begin fun c -> - pf_apply begin fun env sigma -> - try - let (sigma, _) = Typing.type_of env sigma c in - Proofview.Unsafe.tclEVARS sigma >>= fun () -> - return (of_result Value.of_constr (Inl c)) - with e when CErrors.noncritical e -> - let e = CErrors.push e in - return (of_result Value.of_constr (Inr e)) - end -end - -let () = define3 "constr_substnl" (list constr) int constr begin fun subst k c -> - let ans = EConstr.Vars.substnl subst k c in - return (Value.of_constr ans) -end - -let () = define3 "constr_closenl" (list ident) int constr begin fun ids k c -> - let ans = EConstr.Vars.substn_vars k ids c in - return (Value.of_constr ans) -end - -let () = define1 "constr_case" (repr_ext val_inductive) begin fun ind -> - Proofview.tclENV >>= fun env -> - try - let ans = Inductiveops.make_case_info env ind Sorts.Relevant Constr.RegularStyle in - return (Value.of_ext Value.val_case ans) - with e when CErrors.noncritical e -> - throw err_notfound -end - -let () = define2 "constr_constructor" (repr_ext val_inductive) int begin fun (ind, i) k -> - Proofview.tclENV >>= fun env -> - try - let open Declarations in - let ans = Environ.lookup_mind ind env in - let _ = ans.mind_packets.(i).mind_consnames.(k) in - return (Value.of_ext val_constructor ((ind, i), (k + 1))) - with e when CErrors.noncritical e -> - throw err_notfound -end - -let () = define3 "constr_in_context" ident constr closure begin fun id t c -> - Proofview.Goal.goals >>= function - | [gl] -> - gl >>= fun gl -> - let env = Proofview.Goal.env gl in - let sigma = Proofview.Goal.sigma gl in - let has_var = - try - let _ = Environ.lookup_named_val id env in - true - with Not_found -> false - in - if has_var then - Tacticals.New.tclZEROMSG (str "Variable already exists") - else - let open Context.Named.Declaration in - let nenv = EConstr.push_named (LocalAssum (Context.make_annot id Sorts.Relevant, t)) env in - let (sigma, (evt, _)) = Evarutil.new_type_evar nenv sigma Evd.univ_flexible in - let (sigma, evk) = Evarutil.new_pure_evar (Environ.named_context_val nenv) sigma evt in - Proofview.Unsafe.tclEVARS sigma >>= fun () -> - Proofview.Unsafe.tclSETGOALS [Proofview.with_empty_state evk] >>= fun () -> - thaw c >>= fun _ -> - Proofview.Unsafe.tclSETGOALS [Proofview.with_empty_state (Proofview.Goal.goal gl)] >>= fun () -> - let args = List.map (fun d -> EConstr.mkVar (get_id d)) (EConstr.named_context env) in - let args = Array.of_list (EConstr.mkRel 1 :: args) in - let ans = EConstr.mkEvar (evk, args) in - let ans = EConstr.mkLambda (Context.make_annot (Name id) Sorts.Relevant, t, ans) in - return (Value.of_constr ans) - | _ -> - throw err_notfocussed -end - -(** Patterns *) - -let empty_context = EConstr.mkMeta Constr_matching.special_meta - -let () = define0 "pattern_empty_context" begin - return (Value.of_constr empty_context) -end - -let () = define2 "pattern_matches" pattern constr begin fun pat c -> - pf_apply begin fun env sigma -> - let ans = - try Some (Constr_matching.matches env sigma pat c) - with Constr_matching.PatternMatchingFailure -> None - in - begin match ans with - | None -> fail err_matchfailure - | Some ans -> - let ans = Id.Map.bindings ans in - let of_pair (id, c) = Value.of_tuple [| Value.of_ident id; Value.of_constr c |] in - return (Value.of_list of_pair ans) - end - end -end - -let () = define2 "pattern_matches_subterm" pattern constr begin fun pat c -> - let open Constr_matching in - let rec of_ans s = match IStream.peek s with - | IStream.Nil -> fail err_matchfailure - | IStream.Cons ({ m_sub = (_, sub); m_ctx }, s) -> - let ans = Id.Map.bindings sub in - let of_pair (id, c) = Value.of_tuple [| Value.of_ident id; Value.of_constr c |] in - let ans = Value.of_tuple [| Value.of_constr (Lazy.force m_ctx); Value.of_list of_pair ans |] in - Proofview.tclOR (return ans) (fun _ -> of_ans s) - in - pf_apply begin fun env sigma -> - let ans = Constr_matching.match_subterm env sigma (Id.Set.empty,pat) c in - of_ans ans - end -end - -let () = define2 "pattern_matches_vect" pattern constr begin fun pat c -> - pf_apply begin fun env sigma -> - let ans = - try Some (Constr_matching.matches env sigma pat c) - with Constr_matching.PatternMatchingFailure -> None - in - begin match ans with - | None -> fail err_matchfailure - | Some ans -> - let ans = Id.Map.bindings ans in - let ans = Array.map_of_list snd ans in - return (Value.of_array Value.of_constr ans) - end - end -end - -let () = define2 "pattern_matches_subterm_vect" pattern constr begin fun pat c -> - let open Constr_matching in - let rec of_ans s = match IStream.peek s with - | IStream.Nil -> fail err_matchfailure - | IStream.Cons ({ m_sub = (_, sub); m_ctx }, s) -> - let ans = Id.Map.bindings sub in - let ans = Array.map_of_list snd ans in - let ans = Value.of_tuple [| Value.of_constr (Lazy.force m_ctx); Value.of_array Value.of_constr ans |] in - Proofview.tclOR (return ans) (fun _ -> of_ans s) - in - pf_apply begin fun env sigma -> - let ans = Constr_matching.match_subterm env sigma (Id.Set.empty,pat) c in - of_ans ans - end -end - -let () = define3 "pattern_matches_goal" bool (list (pair bool pattern)) (pair bool pattern) begin fun rev hp cp -> - assert_focussed >>= fun () -> - Proofview.Goal.enter_one begin fun gl -> - let env = Proofview.Goal.env gl in - let sigma = Proofview.Goal.sigma gl in - let concl = Proofview.Goal.concl gl in - let mk_pattern (b, pat) = if b then Tac2match.MatchPattern pat else Tac2match.MatchContext pat in - let r = (List.map mk_pattern hp, mk_pattern cp) in - Tac2match.match_goal env sigma concl ~rev r >>= fun (hyps, ctx, subst) -> - let of_ctxopt ctx = Value.of_constr (Option.default empty_context ctx) in - let hids = Value.of_array Value.of_ident (Array.map_of_list fst hyps) in - let hctx = Value.of_array of_ctxopt (Array.map_of_list snd hyps) in - let subs = Value.of_array Value.of_constr (Array.map_of_list snd (Id.Map.bindings subst)) in - let cctx = of_ctxopt ctx in - let ans = Value.of_tuple [| hids; hctx; subs; cctx |] in - Proofview.tclUNIT ans - end -end - -let () = define2 "pattern_instantiate" constr constr begin fun ctx c -> - let ctx = EConstr.Unsafe.to_constr ctx in - let c = EConstr.Unsafe.to_constr c in - let ans = Termops.subst_meta [Constr_matching.special_meta, c] ctx in - return (Value.of_constr (EConstr.of_constr ans)) -end - -(** Error *) - -let () = define1 "throw" exn begin fun (e, info) -> - throw ~info e -end - -(** Control *) - -(** exn -> 'a *) -let () = define1 "zero" exn begin fun (e, info) -> - fail ~info e -end - -(** (unit -> 'a) -> (exn -> 'a) -> 'a *) -let () = define2 "plus" closure closure begin fun x k -> - Proofview.tclOR (thaw x) (fun e -> Tac2ffi.apply k [Value.of_exn e]) -end - -(** (unit -> 'a) -> 'a *) -let () = define1 "once" closure begin fun f -> - Proofview.tclONCE (thaw f) -end - -(** (unit -> unit) list -> unit *) -let () = define1 "dispatch" (list closure) begin fun l -> - let l = List.map (fun f -> Proofview.tclIGNORE (thaw f)) l in - Proofview.tclDISPATCH l >>= fun () -> return v_unit -end - -(** (unit -> unit) list -> (unit -> unit) -> (unit -> unit) list -> unit *) -let () = define3 "extend" (list closure) closure (list closure) begin fun lft tac rgt -> - let lft = List.map (fun f -> Proofview.tclIGNORE (thaw f)) lft in - let tac = Proofview.tclIGNORE (thaw tac) in - let rgt = List.map (fun f -> Proofview.tclIGNORE (thaw f)) rgt in - Proofview.tclEXTEND lft tac rgt >>= fun () -> return v_unit -end - -(** (unit -> unit) -> unit *) -let () = define1 "enter" closure begin fun f -> - let f = Proofview.tclIGNORE (thaw f) in - Proofview.tclINDEPENDENT f >>= fun () -> return v_unit -end - -(** (unit -> 'a) -> ('a * ('exn -> 'a)) result *) -let () = define1 "case" closure begin fun f -> - Proofview.tclCASE (thaw f) >>= begin function - | Proofview.Next (x, k) -> - let k = Tac2ffi.mk_closure arity_one begin fun e -> - let (e, info) = Value.to_exn e in - set_bt info >>= fun info -> - k (e, info) - end in - return (v_blk 0 [| Value.of_tuple [| x; Value.of_closure k |] |]) - | Proofview.Fail e -> return (v_blk 1 [| Value.of_exn e |]) - end -end - -(** int -> int -> (unit -> 'a) -> 'a *) -let () = define3 "focus" int int closure begin fun i j tac -> - Proofview.tclFOCUS i j (thaw tac) -end - -(** unit -> unit *) -let () = define0 "shelve" begin - Proofview.shelve >>= fun () -> return v_unit -end - -(** unit -> unit *) -let () = define0 "shelve_unifiable" begin - Proofview.shelve_unifiable >>= fun () -> return v_unit -end - -let () = define1 "new_goal" int begin fun ev -> - let ev = Evar.unsafe_of_int ev in - Proofview.tclEVARMAP >>= fun sigma -> - if Evd.mem sigma ev then - Proofview.Unsafe.tclNEWGOALS [Proofview.with_empty_state ev] <*> Proofview.tclUNIT v_unit - else throw err_notfound -end - -(** unit -> constr *) -let () = define0 "goal" begin - assert_focussed >>= fun () -> - Proofview.Goal.enter_one begin fun gl -> - let concl = Tacmach.New.pf_nf_concl gl in - return (Value.of_constr concl) - end -end - -(** ident -> constr *) -let () = define1 "hyp" ident begin fun id -> - pf_apply begin fun env _ -> - let mem = try ignore (Environ.lookup_named id env); true with Not_found -> false in - if mem then return (Value.of_constr (EConstr.mkVar id)) - else Tacticals.New.tclZEROMSG - (str "Hypothesis " ++ quote (Id.print id) ++ str " not found") (* FIXME: Do something more sensible *) - end -end - -let () = define0 "hyps" begin - pf_apply begin fun env _ -> - let open Context in - let open Named.Declaration in - let hyps = List.rev (Environ.named_context env) in - let map = function - | LocalAssum (id, t) -> - let t = EConstr.of_constr t in - Value.of_tuple [|Value.of_ident id.binder_name; Value.of_option Value.of_constr None; Value.of_constr t|] - | LocalDef (id, c, t) -> - let c = EConstr.of_constr c in - let t = EConstr.of_constr t in - Value.of_tuple [|Value.of_ident id.binder_name; Value.of_option Value.of_constr (Some c); Value.of_constr t|] - in - return (Value.of_list map hyps) - end -end - -(** (unit -> constr) -> unit *) -let () = define1 "refine" closure begin fun c -> - let c = thaw c >>= fun c -> Proofview.tclUNIT ((), Value.to_constr c) in - Proofview.Goal.enter begin fun gl -> - Refine.generic_refine ~typecheck:true c gl - end >>= fun () -> return v_unit -end - -let () = define2 "with_holes" closure closure begin fun x f -> - Proofview.tclEVARMAP >>= fun sigma0 -> - thaw x >>= fun ans -> - Proofview.tclEVARMAP >>= fun sigma -> - Proofview.Unsafe.tclEVARS sigma0 >>= fun () -> - Tacticals.New.tclWITHHOLES false (Tac2ffi.apply f [ans]) sigma -end - -let () = define1 "progress" closure begin fun f -> - Proofview.tclPROGRESS (thaw f) -end - -let () = define2 "abstract" (option ident) closure begin fun id f -> - Abstract.tclABSTRACT id (Proofview.tclIGNORE (thaw f)) >>= fun () -> - return v_unit -end - -let () = define2 "time" (option string) closure begin fun s f -> - let s = Option.map Bytes.to_string s in - Proofview.tclTIME s (thaw f) -end - -let () = define0 "check_interrupt" begin - Proofview.tclCHECKINTERRUPT >>= fun () -> return v_unit -end - -(** Fresh *) - -let () = define2 "fresh_free_union" (repr_ext val_free) (repr_ext val_free) begin fun set1 set2 -> - let ans = Id.Set.union set1 set2 in - return (Value.of_ext Value.val_free ans) -end - -let () = define1 "fresh_free_of_ids" (list ident) begin fun ids -> - let free = List.fold_right Id.Set.add ids Id.Set.empty in - return (Value.of_ext Value.val_free free) -end - -let () = define1 "fresh_free_of_constr" constr begin fun c -> - Proofview.tclEVARMAP >>= fun sigma -> - let rec fold accu c = match EConstr.kind sigma c with - | Constr.Var id -> Id.Set.add id accu - | _ -> EConstr.fold sigma fold accu c - in - let ans = fold Id.Set.empty c in - return (Value.of_ext Value.val_free ans) -end - -let () = define2 "fresh_fresh" (repr_ext val_free) ident begin fun avoid id -> - let nid = Namegen.next_ident_away_from id (fun id -> Id.Set.mem id avoid) in - return (Value.of_ident nid) -end - -(** Env *) - -let () = define1 "env_get" (list ident) begin fun ids -> - let r = match ids with - | [] -> None - | _ :: _ as ids -> - let (id, path) = List.sep_last ids in - let path = DirPath.make (List.rev path) in - let fp = Libnames.make_path path id in - try Some (Nametab.global_of_path fp) with Not_found -> None - in - return (Value.of_option Value.of_reference r) -end - -let () = define1 "env_expand" (list ident) begin fun ids -> - let r = match ids with - | [] -> [] - | _ :: _ as ids -> - let (id, path) = List.sep_last ids in - let path = DirPath.make (List.rev path) in - let qid = Libnames.make_qualid path id in - Nametab.locate_all qid - in - return (Value.of_list Value.of_reference r) -end - -let () = define1 "env_path" reference begin fun r -> - match Nametab.path_of_global r with - | fp -> - let (path, id) = Libnames.repr_path fp in - let path = DirPath.repr path in - return (Value.of_list Value.of_ident (List.rev_append path [id])) - | exception Not_found -> - throw err_notfound -end - -let () = define1 "env_instantiate" reference begin fun r -> - Proofview.tclENV >>= fun env -> - Proofview.tclEVARMAP >>= fun sigma -> - let (sigma, c) = Evd.fresh_global env sigma r in - Proofview.Unsafe.tclEVARS sigma >>= fun () -> - return (Value.of_constr c) -end - -(** Ltac1 in Ltac2 *) - -let ltac1 = Tac2ffi.repr_ext Value.val_ltac1 -let of_ltac1 v = Value.of_ext Value.val_ltac1 v - -let () = define1 "ltac1_ref" (list ident) begin fun ids -> - let open Ltac_plugin in - let r = match ids with - | [] -> raise Not_found - | _ :: _ as ids -> - let (id, path) = List.sep_last ids in - let path = DirPath.make (List.rev path) in - let fp = Libnames.make_path path id in - if Tacenv.exists_tactic fp then - List.hd (Tacenv.locate_extended_all_tactic (Libnames.qualid_of_path fp)) - else raise Not_found - in - let tac = Tacinterp.Value.of_closure (Tacinterp.default_ist ()) (Tacenv.interp_ltac r) in - return (Value.of_ext val_ltac1 tac) -end - -let () = define1 "ltac1_run" ltac1 begin fun v -> - let open Ltac_plugin in - Tacinterp.tactic_of_value (Tacinterp.default_ist ()) v >>= fun () -> - return v_unit -end - -let () = define3 "ltac1_apply" ltac1 (list ltac1) closure begin fun f args k -> - let open Ltac_plugin in - let open Tacexpr in - let open Locus in - let k ret = - Proofview.tclIGNORE (Tac2ffi.apply k [Value.of_ext val_ltac1 ret]) - in - let fold arg (i, vars, lfun) = - let id = Id.of_string ("x" ^ string_of_int i) in - let x = Reference (ArgVar CAst.(make 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 - let tac = TacArg(CAst.make @@ TacCall (CAst.make (ArgVar CAst.(make @@ Id.of_string "F"),args))) in - Tacinterp.val_interp ist tac k >>= fun () -> - return v_unit -end - -let () = define1 "ltac1_of_constr" constr begin fun c -> - let open Ltac_plugin in - return (Value.of_ext val_ltac1 (Tacinterp.Value.of_constr c)) -end - -let () = define1 "ltac1_to_constr" ltac1 begin fun v -> - let open Ltac_plugin in - return (Value.of_option Value.of_constr (Tacinterp.Value.to_constr v)) -end - -let () = define1 "ltac1_of_list" (list ltac1) begin fun l -> - let open Geninterp.Val in - return (Value.of_ext val_ltac1 (inject (Base typ_list) l)) -end - -let () = define1 "ltac1_to_list" ltac1 begin fun v -> - let open Ltac_plugin in - return (Value.of_option (Value.of_list of_ltac1) (Tacinterp.Value.to_list v)) -end - -(** ML types *) - -let constr_flags () = - let open Pretyping in - { - use_typeclasses = true; - solve_unification_constraints = true; - fail_evar = true; - expand_evars = true; - program_mode = false; - polymorphic = false; - } - -let open_constr_no_classes_flags () = - let open Pretyping in - { - use_typeclasses = false; - solve_unification_constraints = true; - fail_evar = false; - expand_evars = true; - program_mode = false; - polymorphic = false; - } - -(** Embed all Ltac2 data into Values *) -let to_lvar ist = - let open Glob_ops in - let lfun = Tac2interp.set_env ist Id.Map.empty in - { empty_lvar with Ltac_pretype.ltac_genargs = lfun } - -let gtypref kn = GTypRef (Other kn, []) - -let intern_constr self ist c = - let (_, (c, _)) = Genintern.intern Stdarg.wit_constr ist c in - (GlbVal c, gtypref t_constr) - -let catchable_exception = function - | Logic_monad.Exception _ -> false - | e -> CErrors.noncritical e - -let interp_constr flags ist c = - let open Pretyping in - let ist = to_lvar ist in - pf_apply begin fun env sigma -> - try - let (sigma, c) = understand_ltac flags env sigma ist WithoutTypeConstraint c in - let c = Value.of_constr c in - Proofview.Unsafe.tclEVARS sigma >>= fun () -> - Proofview.tclUNIT c - with e when catchable_exception e -> - let (e, info) = CErrors.push e in - set_bt info >>= fun info -> - match Exninfo.get info fatal_flag with - | None -> Proofview.tclZERO ~info e - | Some () -> throw ~info e - end - -let () = - let intern = intern_constr in - let interp ist c = interp_constr (constr_flags ()) ist c in - let print env c = str "constr:(" ++ Printer.pr_lglob_constr_env env c ++ str ")" in - let subst subst c = Detyping.subst_glob_constr (Global.env()) subst c in - let obj = { - ml_intern = intern; - ml_subst = subst; - ml_interp = interp; - ml_print = print; - } in - define_ml_object Tac2quote.wit_constr obj - -let () = - let intern = intern_constr in - let interp ist c = interp_constr (open_constr_no_classes_flags ()) ist c in - let print env c = str "open_constr:(" ++ Printer.pr_lglob_constr_env env c ++ str ")" in - let subst subst c = Detyping.subst_glob_constr (Global.env()) subst c in - let obj = { - ml_intern = intern; - ml_subst = subst; - ml_interp = interp; - ml_print = print; - } in - define_ml_object Tac2quote.wit_open_constr obj - -let () = - let interp _ id = return (Value.of_ident id) in - let print _ id = str "ident:(" ++ Id.print id ++ str ")" in - let obj = { - ml_intern = (fun _ _ id -> GlbVal id, gtypref t_ident); - ml_interp = interp; - ml_subst = (fun _ id -> id); - ml_print = print; - } in - define_ml_object Tac2quote.wit_ident obj - -let () = - let intern self ist c = - let env = ist.Genintern.genv in - let sigma = Evd.from_env env in - let warn = if !Ltac_plugin.Tacintern.strict_check then fun x -> x else Constrintern.for_grammar in - let _, pat = warn (fun () ->Constrintern.intern_constr_pattern env sigma ~as_type:false c) () in - GlbVal pat, gtypref t_pattern - in - let subst subst c = - let env = Global.env () in - let sigma = Evd.from_env env in - Patternops.subst_pattern env sigma subst c - in - let print env pat = str "pattern:(" ++ Printer.pr_lconstr_pattern_env env Evd.empty pat ++ str ")" in - let interp _ c = return (Value.of_pattern c) in - let obj = { - ml_intern = intern; - ml_interp = interp; - ml_subst = subst; - ml_print = print; - } in - define_ml_object Tac2quote.wit_pattern obj - -let () = - let intern self ist ref = match ref.CAst.v with - | Tac2qexpr.QHypothesis id -> - GlbVal (Globnames.VarRef id), gtypref t_reference - | Tac2qexpr.QReference qid -> - let gr = - try Nametab.locate qid - with Not_found -> - Nametab.error_global_not_found qid - in - GlbVal gr, gtypref t_reference - in - let subst s c = Globnames.subst_global_reference s c in - let interp _ gr = return (Value.of_reference gr) in - let print _ = function - | Globnames.VarRef id -> str "reference:(" ++ str "&" ++ Id.print id ++ str ")" - | r -> str "reference:(" ++ Printer.pr_global r ++ str ")" - in - let obj = { - ml_intern = intern; - ml_subst = subst; - ml_interp = interp; - ml_print = print; - } in - define_ml_object Tac2quote.wit_reference obj - -let () = - let intern self ist tac = - (* Prevent inner calls to Ltac2 values *) - let extra = Tac2intern.drop_ltac2_env ist.Genintern.extra in - let ist = { ist with Genintern.extra } in - let _, tac = Genintern.intern Ltac_plugin.Tacarg.wit_tactic ist tac in - GlbVal tac, gtypref t_unit - in - let interp ist tac = - let ist = { env_ist = Id.Map.empty } in - let lfun = Tac2interp.set_env ist Id.Map.empty in - let ist = Ltac_plugin.Tacinterp.default_ist () in - let ist = { ist with Geninterp.lfun = lfun } in - let tac = (Ltac_plugin.Tacinterp.eval_tactic_ist ist tac : unit Proofview.tactic) in - let wrap (e, info) = set_bt info >>= fun info -> Proofview.tclZERO ~info e in - Proofview.tclOR tac wrap >>= fun () -> - return v_unit - in - let subst s tac = Genintern.substitute Ltac_plugin.Tacarg.wit_tactic s tac in - let print env tac = - str "ltac1:(" ++ Ltac_plugin.Pptactic.pr_glob_tactic env tac ++ str ")" - in - let obj = { - ml_intern = intern; - ml_subst = subst; - ml_interp = interp; - ml_print = print; - } in - define_ml_object Tac2quote.wit_ltac1 obj - -let () = - let open Ltac_plugin in - let intern self ist tac = - (* Prevent inner calls to Ltac2 values *) - let extra = Tac2intern.drop_ltac2_env ist.Genintern.extra in - let ist = { ist with Genintern.extra } in - let _, tac = Genintern.intern Ltac_plugin.Tacarg.wit_tactic ist tac in - GlbVal tac, gtypref t_ltac1 - in - let interp ist tac = - let ist = { env_ist = Id.Map.empty } in - let lfun = Tac2interp.set_env ist Id.Map.empty in - let ist = Ltac_plugin.Tacinterp.default_ist () in - let ist = { ist with Geninterp.lfun = lfun } in - return (Value.of_ext val_ltac1 (Tacinterp.Value.of_closure ist tac)) - in - let subst s tac = Genintern.substitute Tacarg.wit_tactic s tac in - let print env tac = - str "ltac1val:(" ++ Ltac_plugin.Pptactic.pr_glob_tactic env tac ++ str ")" - in - let obj = { - ml_intern = intern; - ml_subst = subst; - ml_interp = interp; - ml_print = print; - } in - define_ml_object Tac2quote.wit_ltac1val obj - -(** Ltac2 in terms *) - -let () = - let interp ist poly env sigma concl tac = - let ist = Tac2interp.get_env ist in - let tac = Proofview.tclIGNORE (Tac2interp.interp ist tac) in - let name, poly = Id.of_string "ltac2", poly in - let c, sigma = Pfedit.refine_by_tactic ~name ~poly env sigma concl tac in - (EConstr.of_constr c, sigma) - in - GlobEnv.register_constr_interp0 wit_ltac2 interp - -let () = - let interp ist poly env sigma concl id = - let ist = Tac2interp.get_env ist in - let c = Id.Map.find id ist.env_ist in - let c = Value.to_constr c in - let sigma = Typing.check env sigma c concl in - (c, sigma) - in - GlobEnv.register_constr_interp0 wit_ltac2_quotation interp - -let () = - let pr_raw id = Genprint.PrinterBasic (fun _env _sigma -> mt ()) in - let pr_glb id = Genprint.PrinterBasic (fun _env _sigma -> str "$" ++ Id.print id) in - let pr_top _ = Genprint.TopPrinterBasic mt in - Genprint.register_print0 wit_ltac2_quotation pr_raw pr_glb pr_top - -(** Ltac2 in Ltac1 *) - -let () = - let e = Tac2entries.Pltac.tac2expr in - let inject (loc, v) = Ltac_plugin.Tacexpr.TacGeneric (in_gen (rawwit wit_ltac2) v) in - Ltac_plugin.Tacentries.create_ltac_quotation "ltac2" inject (e, None) - -let () = - let open Ltac_plugin in - let open Tacinterp in - let idtac = Value.of_closure (default_ist ()) (Tacexpr.TacId []) in - let interp ist tac = -(* let ist = Tac2interp.get_env ist.Geninterp.lfun in *) - let ist = { env_ist = Id.Map.empty } in - Tac2interp.interp ist tac >>= fun _ -> - Ftactic.return idtac - in - Geninterp.register_interp0 wit_ltac2 interp - -let () = - let pr_raw _ = Genprint.PrinterBasic (fun _env _sigma -> mt ()) in - let pr_glb e = Genprint.PrinterBasic (fun _env _sigma -> Tac2print.pr_glbexpr e) in - let pr_top _ = Genprint.TopPrinterBasic mt in - Genprint.register_print0 wit_ltac2 pr_raw pr_glb pr_top - -(** Built-in notation scopes *) - -let add_scope s f = - Tac2entries.register_scope (Id.of_string s) f - -let rec pr_scope = let open CAst in function -| SexprStr {v=s} -> qstring s -| SexprInt {v=n} -> Pp.int n -| SexprRec (_, {v=na}, args) -> - let na = match na with - | None -> str "_" - | Some id -> Id.print id - in - na ++ str "(" ++ prlist_with_sep (fun () -> str ", ") pr_scope args ++ str ")" - -let scope_fail s args = - let args = str "(" ++ prlist_with_sep (fun () -> str ", ") pr_scope args ++ str ")" in - CErrors.user_err (str "Invalid arguments " ++ args ++ str " in scope " ++ str s) - -let q_unit = CAst.make @@ CTacCst (AbsKn (Tuple 0)) - -let add_generic_scope s entry arg = - let parse = function - | [] -> - let scope = Extend.Aentry entry in - let act x = CAst.make @@ CTacExt (arg, x) in - Tac2entries.ScopeRule (scope, act) - | arg -> scope_fail s arg - in - add_scope s parse - -open CAst - -let () = add_scope "keyword" begin function -| [SexprStr {loc;v=s}] -> - let scope = Extend.Atoken (Tok.PKEYWORD s) in - Tac2entries.ScopeRule (scope, (fun _ -> q_unit)) -| arg -> scope_fail "keyword" arg -end - -let () = add_scope "terminal" begin function -| [SexprStr {loc;v=s}] -> - let scope = Extend.Atoken (CLexer.terminal s) in - Tac2entries.ScopeRule (scope, (fun _ -> q_unit)) -| arg -> scope_fail "terminal" arg -end - -let () = add_scope "list0" begin function -| [tok] -> - let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in - let scope = Extend.Alist0 scope in - let act l = Tac2quote.of_list act l in - Tac2entries.ScopeRule (scope, act) -| [tok; SexprStr {v=str}] -> - let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in - let sep = Extend.Atoken (CLexer.terminal str) in - let scope = Extend.Alist0sep (scope, sep) in - let act l = Tac2quote.of_list act l in - Tac2entries.ScopeRule (scope, act) -| arg -> scope_fail "list0" arg -end - -let () = add_scope "list1" begin function -| [tok] -> - let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in - let scope = Extend.Alist1 scope in - let act l = Tac2quote.of_list act l in - Tac2entries.ScopeRule (scope, act) -| [tok; SexprStr {v=str}] -> - let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in - let sep = Extend.Atoken (CLexer.terminal str) in - let scope = Extend.Alist1sep (scope, sep) in - let act l = Tac2quote.of_list act l in - Tac2entries.ScopeRule (scope, act) -| arg -> scope_fail "list1" arg -end - -let () = add_scope "opt" begin function -| [tok] -> - let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in - let scope = Extend.Aopt scope in - let act opt = match opt with - | None -> - CAst.make @@ CTacCst (AbsKn (Other Core.c_none)) - | Some x -> - CAst.make @@ CTacApp (CAst.make @@ CTacCst (AbsKn (Other Core.c_some)), [act x]) - in - Tac2entries.ScopeRule (scope, act) -| arg -> scope_fail "opt" arg -end - -let () = add_scope "self" begin function -| [] -> - let scope = Extend.Aself in - let act tac = tac in - Tac2entries.ScopeRule (scope, act) -| arg -> scope_fail "self" arg -end - -let () = add_scope "next" begin function -| [] -> - let scope = Extend.Anext in - let act tac = tac in - Tac2entries.ScopeRule (scope, act) -| arg -> scope_fail "next" arg -end - -let () = add_scope "tactic" begin function -| [] -> - (* Default to level 5 parsing *) - let scope = Extend.Aentryl (tac2expr, "5") in - let act tac = tac in - Tac2entries.ScopeRule (scope, act) -| [SexprInt {loc;v=n}] as arg -> - let () = if n < 0 || n > 6 then scope_fail "tactic" arg in - let scope = Extend.Aentryl (tac2expr, string_of_int n) in - let act tac = tac in - Tac2entries.ScopeRule (scope, act) -| arg -> scope_fail "tactic" arg -end - -let () = add_scope "thunk" begin function -| [tok] -> - let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in - let act e = Tac2quote.thunk (act e) in - Tac2entries.ScopeRule (scope, act) -| arg -> scope_fail "thunk" arg -end - -let add_expr_scope name entry f = - add_scope name begin function - | [] -> Tac2entries.ScopeRule (Extend.Aentry entry, f) - | arg -> scope_fail name arg - end - -let () = add_expr_scope "ident" q_ident (fun id -> Tac2quote.of_anti Tac2quote.of_ident id) -let () = add_expr_scope "bindings" q_bindings Tac2quote.of_bindings -let () = add_expr_scope "with_bindings" q_with_bindings Tac2quote.of_bindings -let () = add_expr_scope "intropattern" q_intropattern Tac2quote.of_intro_pattern -let () = add_expr_scope "intropatterns" q_intropatterns Tac2quote.of_intro_patterns -let () = add_expr_scope "destruction_arg" q_destruction_arg Tac2quote.of_destruction_arg -let () = add_expr_scope "induction_clause" q_induction_clause Tac2quote.of_induction_clause -let () = add_expr_scope "conversion" q_conversion Tac2quote.of_conversion -let () = add_expr_scope "rewriting" q_rewriting Tac2quote.of_rewriting -let () = add_expr_scope "clause" q_clause Tac2quote.of_clause -let () = add_expr_scope "hintdb" q_hintdb Tac2quote.of_hintdb -let () = add_expr_scope "occurrences" q_occurrences Tac2quote.of_occurrences -let () = add_expr_scope "dispatch" q_dispatch Tac2quote.of_dispatch -let () = add_expr_scope "strategy" q_strategy_flag Tac2quote.of_strategy_flag -let () = add_expr_scope "reference" q_reference Tac2quote.of_reference -let () = add_expr_scope "move_location" q_move_location Tac2quote.of_move_location -let () = add_expr_scope "pose" q_pose Tac2quote.of_pose -let () = add_expr_scope "assert" q_assert Tac2quote.of_assertion -let () = add_expr_scope "constr_matching" q_constr_matching Tac2quote.of_constr_matching -let () = add_expr_scope "goal_matching" q_goal_matching Tac2quote.of_goal_matching - -let () = add_generic_scope "constr" Pcoq.Constr.constr Tac2quote.wit_constr -let () = add_generic_scope "open_constr" Pcoq.Constr.constr Tac2quote.wit_open_constr -let () = add_generic_scope "pattern" Pcoq.Constr.constr Tac2quote.wit_pattern - -(** seq scope, a bit hairy *) - -open Extend -exception SelfSymbol - -let rec generalize_symbol : - type a tr s. (s, tr, a) Extend.symbol -> (s, Extend.norec, a) Extend.symbol = function -| Atoken tok -> Atoken tok -| Alist1 e -> Alist1 (generalize_symbol e) -| Alist1sep (e, sep) -> - let e = generalize_symbol e in - let sep = generalize_symbol sep in - Alist1sep (e, sep) -| Alist0 e -> Alist0 (generalize_symbol e) -| Alist0sep (e, sep) -> - let e = generalize_symbol e in - let sep = generalize_symbol sep in - Alist0sep (e, sep) -| Aopt e -> Aopt (generalize_symbol e) -| Aself -> raise SelfSymbol -| Anext -> raise SelfSymbol -| Aentry e -> Aentry e -| Aentryl (e, l) -> Aentryl (e, l) -| Arules r -> Arules r - -type _ converter = -| CvNil : (Loc.t -> raw_tacexpr) converter -| CvCns : 'act converter * ('a -> raw_tacexpr) option -> ('a -> 'act) converter - -let rec apply : type a. a converter -> raw_tacexpr list -> a = function -| CvNil -> fun accu loc -> Tac2quote.of_tuple ~loc accu -| CvCns (c, None) -> fun accu x -> apply c accu -| CvCns (c, Some f) -> fun accu x -> apply c (f x :: accu) - -type seqrule = -| Seqrule : (Tac2expr.raw_tacexpr, Extend.norec, 'act, Loc.t -> raw_tacexpr) rule * 'act converter -> seqrule - -let rec make_seq_rule = function -| [] -> - Seqrule (Stop, CvNil) -| tok :: rem -> - let Tac2entries.ScopeRule (scope, f) = Tac2entries.parse_scope tok in - let scope = generalize_symbol scope in - let Seqrule (r, c) = make_seq_rule rem in - let r = NextNoRec (r, scope) in - let f = match tok with - | SexprStr _ -> None (* Leave out mere strings *) - | _ -> Some f - in - Seqrule (r, CvCns (c, f)) - -let () = add_scope "seq" begin fun toks -> - let scope = - try - let Seqrule (r, c) = make_seq_rule (List.rev toks) in - Arules [Rules (r, apply c [])] - with SelfSymbol -> - CErrors.user_err (str "Recursive symbols (self / next) are not allowed in local rules") - in - Tac2entries.ScopeRule (scope, (fun e -> e)) -end diff --git a/src/tac2core.mli b/src/tac2core.mli deleted file mode 100644 index 9fae65bb3e..0000000000 --- a/src/tac2core.mli +++ /dev/null @@ -1,30 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* Evd.evar_map -> 'a Proofview.tactic) -> 'a Proofview.tactic diff --git a/src/tac2dyn.ml b/src/tac2dyn.ml deleted file mode 100644 index 896676f08b..0000000000 --- a/src/tac2dyn.ml +++ /dev/null @@ -1,27 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* ('raw * 'glb) pack - include Arg.Map(struct type 'a t = 'a pack end) -end - -module Val = Dyn.Make(struct end) diff --git a/src/tac2dyn.mli b/src/tac2dyn.mli deleted file mode 100644 index e995296840..0000000000 --- a/src/tac2dyn.mli +++ /dev/null @@ -1,34 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* ('a, 'b) tag - val eq : ('a1, 'b1) tag -> ('a2, 'b2) tag -> ('a1 * 'b1, 'a2 * 'b2) CSig.eq option - val repr : ('a, 'b) tag -> string -end -(** Arguments that are part of an AST. *) - -module type Param = sig type ('raw, 'glb) t end - -module ArgMap (M : Param) : -sig - type _ pack = Pack : ('raw, 'glb) M.t -> ('raw * 'glb) pack - type t - val empty : t - val add : ('a, 'b) Arg.tag -> ('a * 'b) pack -> t -> t - val remove : ('a, 'b) Arg.tag -> t -> t - val find : ('a, 'b) Arg.tag -> t -> ('a * 'b) pack - val mem : ('a, 'b) Arg.tag -> t -> bool -end - -module Val : Dyn.S -(** Toplevel values *) diff --git a/src/tac2entries.ml b/src/tac2entries.ml deleted file mode 100644 index 9fd01426de..0000000000 --- a/src/tac2entries.ml +++ /dev/null @@ -1,938 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* obj = - declare_object {(default_object "TAC2-DEFINITION") with - cache_function = cache_tacdef; - load_function = load_tacdef; - open_function = open_tacdef; - subst_function = subst_tacdef; - classify_function = classify_tacdef} - -(** Type definition *) - -type typdef = { - typdef_local : bool; - typdef_expr : glb_quant_typedef; -} - -let change_kn_label kn id = - let mp = KerName.modpath kn in - KerName.make mp (Label.of_id id) - -let change_sp_label sp id = - let (dp, _) = Libnames.repr_path sp in - Libnames.make_path dp id - -let push_typedef visibility sp kn (_, def) = match def with -| GTydDef _ -> - Tac2env.push_type visibility sp kn -| GTydAlg { galg_constructors = cstrs } -> - (* Register constructors *) - let iter (c, _) = - let spc = change_sp_label sp c in - let knc = change_kn_label kn c in - Tac2env.push_constructor visibility spc knc - in - Tac2env.push_type visibility sp kn; - List.iter iter cstrs -| GTydRec fields -> - (* Register fields *) - let iter (c, _, _) = - let spc = change_sp_label sp c in - let knc = change_kn_label kn c in - Tac2env.push_projection visibility spc knc - in - Tac2env.push_type visibility sp kn; - List.iter iter fields -| GTydOpn -> - Tac2env.push_type visibility sp kn - -let next i = - let ans = !i in - let () = incr i in - ans - -let define_typedef kn (params, def as qdef) = match def with -| GTydDef _ -> - Tac2env.define_type kn qdef -| GTydAlg { galg_constructors = cstrs } -> - (* Define constructors *) - let constant = ref 0 in - let nonconstant = ref 0 in - let iter (c, args) = - let knc = change_kn_label kn c in - let tag = if List.is_empty args then next constant else next nonconstant in - let data = { - Tac2env.cdata_prms = params; - cdata_type = kn; - cdata_args = args; - cdata_indx = Some tag; - } in - Tac2env.define_constructor knc data - in - Tac2env.define_type kn qdef; - List.iter iter cstrs -| GTydRec fs -> - (* Define projections *) - let iter i (id, mut, t) = - let knp = change_kn_label kn id in - let proj = { - Tac2env.pdata_prms = params; - pdata_type = kn; - pdata_ptyp = t; - pdata_mutb = mut; - pdata_indx = i; - } in - Tac2env.define_projection knp proj - in - Tac2env.define_type kn qdef; - List.iteri iter fs -| GTydOpn -> - Tac2env.define_type kn qdef - -let perform_typdef vs ((sp, kn), def) = - let () = if not def.typdef_local then push_typedef vs sp kn def.typdef_expr in - define_typedef kn def.typdef_expr - -let load_typdef i obj = perform_typdef (Until i) obj -let open_typdef i obj = perform_typdef (Exactly i) obj - -let cache_typdef ((sp, kn), def) = - let () = push_typedef (Until 1) sp kn def.typdef_expr in - define_typedef kn def.typdef_expr - -let subst_typdef (subst, def) = - let expr' = subst_quant_typedef subst def.typdef_expr in - if expr' == def.typdef_expr then def else { def with typdef_expr = expr' } - -let classify_typdef o = Substitute o - -let inTypDef : typdef -> obj = - declare_object {(default_object "TAC2-TYPE-DEFINITION") with - cache_function = cache_typdef; - load_function = load_typdef; - open_function = open_typdef; - subst_function = subst_typdef; - classify_function = classify_typdef} - -(** Type extension *) - -type extension_data = { - edata_name : Id.t; - edata_args : int glb_typexpr list; -} - -type typext = { - typext_local : bool; - typext_prms : int; - typext_type : type_constant; - typext_expr : extension_data list; -} - -let push_typext vis sp kn def = - let iter data = - let spc = change_sp_label sp data.edata_name in - let knc = change_kn_label kn data.edata_name in - Tac2env.push_constructor vis spc knc - in - List.iter iter def.typext_expr - -let define_typext kn def = - let iter data = - let knc = change_kn_label kn data.edata_name in - let cdata = { - Tac2env.cdata_prms = def.typext_prms; - cdata_type = def.typext_type; - cdata_args = data.edata_args; - cdata_indx = None; - } in - Tac2env.define_constructor knc cdata - in - List.iter iter def.typext_expr - -let cache_typext ((sp, kn), def) = - let () = define_typext kn def in - push_typext (Until 1) sp kn def - -let perform_typext vs ((sp, kn), def) = - let () = if not def.typext_local then push_typext vs sp kn def in - define_typext kn def - -let load_typext i obj = perform_typext (Until i) obj -let open_typext i obj = perform_typext (Exactly i) obj - -let subst_typext (subst, e) = - let open Mod_subst in - let subst_data data = - let edata_args = List.Smart.map (fun e -> subst_type subst e) data.edata_args in - if edata_args == data.edata_args then data - else { data with edata_args } - in - let typext_type = subst_kn subst e.typext_type in - let typext_expr = List.Smart.map subst_data e.typext_expr in - if typext_type == e.typext_type && typext_expr == e.typext_expr then - e - else - { e with typext_type; typext_expr } - -let classify_typext o = Substitute o - -let inTypExt : typext -> obj = - declare_object {(default_object "TAC2-TYPE-EXTENSION") with - cache_function = cache_typext; - load_function = load_typext; - open_function = open_typext; - subst_function = subst_typext; - classify_function = classify_typext} - -(** Toplevel entries *) - -let fresh_var avoid x = - let bad id = - Id.Set.mem id avoid || - (try ignore (Tac2env.locate_ltac (qualid_of_ident id)); true with Not_found -> false) - in - Namegen.next_ident_away_from (Id.of_string x) bad - -let extract_pattern_type ({loc;v=p} as pat) = match p with -| CPatCnv (pat, ty) -> pat, Some ty -| CPatVar _ | CPatRef _ -> pat, None - -(** Mangle recursive tactics *) -let inline_rec_tactic tactics = - let avoid = List.fold_left (fun accu ({v=id}, _) -> Id.Set.add id accu) Id.Set.empty tactics in - let map (id, e) = match e.v with - | CTacFun (pat, _) -> (id, List.map extract_pattern_type pat, e) - | _ -> - user_err ?loc:id.loc (str "Recursive tactic definitions must be functions") - in - let tactics = List.map map tactics in - let map (id, pat, e) = - let fold_var (avoid, ans) (pat, _) = - let id = fresh_var avoid "x" in - let loc = pat.loc in - (Id.Set.add id avoid, CAst.make ?loc id :: ans) - in - (* Fresh variables to abstract over the function patterns *) - let _, vars = List.fold_left fold_var (avoid, []) pat in - let map_body ({loc;v=id}, _, e) = CAst.(make ?loc @@ CPatVar (Name id)), e in - let bnd = List.map map_body tactics in - let pat_of_id {loc;v=id} = CAst.make ?loc @@ CPatVar (Name id) in - let var_of_id {loc;v=id} = - let qid = qualid_of_ident ?loc id in - CAst.make ?loc @@ CTacRef (RelId qid) - in - let loc0 = e.loc in - let vpat = List.map pat_of_id vars in - let varg = List.map var_of_id vars in - let e = CAst.make ?loc:loc0 @@ CTacLet (true, bnd, CAst.make ?loc:loc0 @@ CTacApp (var_of_id id, varg)) in - (id, CAst.make ?loc:loc0 @@ CTacFun (vpat, e)) - in - List.map map tactics - -let check_lowercase {loc;v=id} = - if Tac2env.is_constructor (Libnames.qualid_of_ident id) then - user_err ?loc (str "The identifier " ++ Id.print id ++ str " must be lowercase") - -let register_ltac ?(local = false) ?(mut = false) isrec tactics = - let map ({loc;v=na}, e) = - let id = match na with - | Anonymous -> - user_err ?loc (str "Tactic definition must have a name") - | Name id -> id - in - let () = check_lowercase CAst.(make ?loc id) in - (CAst.(make ?loc id), e) - in - let tactics = List.map map tactics in - let tactics = - if isrec then inline_rec_tactic tactics else tactics - in - let map ({loc;v=id}, e) = - let (e, t) = intern ~strict:true e in - let () = - if not (is_value e) then - user_err ?loc (str "Tactic definition must be a syntactical value") - in - let kn = Lib.make_kn id in - let exists = - try let _ = Tac2env.interp_global kn in true with Not_found -> false - in - let () = - if exists then - user_err ?loc (str "Tactic " ++ Names.Id.print id ++ str " already exists") - in - (id, e, t) - in - let defs = List.map map tactics in - let iter (id, e, t) = - let def = { - tacdef_local = local; - tacdef_mutable = mut; - tacdef_expr = e; - tacdef_type = t; - } in - ignore (Lib.add_leaf id (inTacDef def)) - in - List.iter iter defs - -let qualid_to_ident qid = - if qualid_is_ident qid then CAst.make ?loc:qid.CAst.loc @@ qualid_basename qid - else user_err ?loc:qid.CAst.loc (str "Identifier expected") - -let register_typedef ?(local = false) isrec types = - let same_name ({v=id1}, _) ({v=id2}, _) = Id.equal id1 id2 in - let () = match List.duplicates same_name types with - | [] -> () - | ({loc;v=id}, _) :: _ -> - user_err ?loc (str "Multiple definition of the type name " ++ Id.print id) - in - let check ({loc;v=id}, (params, def)) = - let same_name {v=id1} {v=id2} = Id.equal id1 id2 in - let () = match List.duplicates same_name params with - | [] -> () - | {loc;v=id} :: _ -> - user_err ?loc (str "The type parameter " ++ Id.print id ++ - str " occurs several times") - in - match def with - | CTydDef _ -> - if isrec then - user_err ?loc (str "The type abbreviation " ++ Id.print id ++ - str " cannot be recursive") - | CTydAlg cs -> - let same_name (id1, _) (id2, _) = Id.equal id1 id2 in - let () = match List.duplicates same_name cs with - | [] -> () - | (id, _) :: _ -> - user_err (str "Multiple definitions of the constructor " ++ Id.print id) - in - () - | CTydRec ps -> - let same_name (id1, _, _) (id2, _, _) = Id.equal id1 id2 in - let () = match List.duplicates same_name ps with - | [] -> () - | (id, _, _) :: _ -> - user_err (str "Multiple definitions of the projection " ++ Id.print id) - in - () - | CTydOpn -> - if isrec then - user_err ?loc (str "The open type declaration " ++ Id.print id ++ - str " cannot be recursive") - in - let () = List.iter check types in - let self = - if isrec then - let fold accu ({v=id}, (params, _)) = - Id.Map.add id (Lib.make_kn id, List.length params) accu - in - List.fold_left fold Id.Map.empty types - else Id.Map.empty - in - let map ({v=id}, def) = - let typdef = { - typdef_local = local; - typdef_expr = intern_typedef self def; - } in - (id, typdef) - in - let types = List.map map types in - let iter (id, def) = ignore (Lib.add_leaf id (inTypDef def)) in - List.iter iter types - -let register_primitive ?(local = false) {loc;v=id} t ml = - let t = intern_open_type t in - let rec count_arrow = function - | GTypArrow (_, t) -> 1 + count_arrow t - | _ -> 0 - in - let arrows = count_arrow (snd t) in - let () = if Int.equal arrows 0 then - user_err ?loc (str "External tactic must have at least one argument") in - let () = - try let _ = Tac2env.interp_primitive ml in () with Not_found -> - user_err ?loc (str "Unregistered primitive " ++ - quote (str ml.mltac_plugin) ++ spc () ++ quote (str ml.mltac_tactic)) - in - let init i = Id.of_string (Printf.sprintf "x%i" i) in - let names = List.init arrows init in - let bnd = List.map (fun id -> Name id) names in - let arg = List.map (fun id -> GTacVar id) names in - let e = GTacFun (bnd, GTacPrm (ml, arg)) in - let def = { - tacdef_local = local; - tacdef_mutable = false; - tacdef_expr = e; - tacdef_type = t; - } in - ignore (Lib.add_leaf id (inTacDef def)) - -let register_open ?(local = false) qid (params, def) = - let kn = - try Tac2env.locate_type qid - with Not_found -> - user_err ?loc:qid.CAst.loc (str "Unbound type " ++ pr_qualid qid) - in - let (tparams, t) = Tac2env.interp_type kn in - let () = match t with - | GTydOpn -> () - | GTydAlg _ | GTydRec _ | GTydDef _ -> - user_err ?loc:qid.CAst.loc (str "Type " ++ pr_qualid qid ++ str " is not an open type") - in - let () = - if not (Int.equal (List.length params) tparams) then - Tac2intern.error_nparams_mismatch ?loc:qid.CAst.loc (List.length params) tparams - in - match def with - | CTydOpn -> () - | CTydAlg def -> - let intern_type t = - let tpe = CTydDef (Some t) in - let (_, ans) = intern_typedef Id.Map.empty (params, tpe) in - match ans with - | GTydDef (Some t) -> t - | _ -> assert false - in - let map (id, tpe) = - let tpe = List.map intern_type tpe in - { edata_name = id; edata_args = tpe } - in - let def = List.map map def in - let def = { - typext_local = local; - typext_type = kn; - typext_prms = tparams; - typext_expr = def; - } in - Lib.add_anonymous_leaf (inTypExt def) - | CTydRec _ | CTydDef _ -> - user_err ?loc:qid.CAst.loc (str "Extensions only accept inductive constructors") - -let register_type ?local isrec types = match types with -| [qid, true, def] -> - let () = if isrec then user_err ?loc:qid.CAst.loc (str "Extensions cannot be recursive") in - register_open ?local qid def -| _ -> - let map (qid, redef, def) = - let () = if redef then - user_err ?loc:qid.loc (str "Types can only be extended one by one") - in - (qualid_to_ident qid, def) - in - let types = List.map map types in - register_typedef ?local isrec types - -(** Parsing *) - -type 'a token = -| TacTerm of string -| TacNonTerm of Name.t * 'a - -type scope_rule = -| ScopeRule : (raw_tacexpr, _, 'a) Extend.symbol * ('a -> raw_tacexpr) -> scope_rule - -type scope_interpretation = sexpr list -> scope_rule - -let scope_table : scope_interpretation Id.Map.t ref = ref Id.Map.empty - -let register_scope id s = - scope_table := Id.Map.add id s !scope_table - -module ParseToken = -struct - -let loc_of_token = function -| SexprStr {loc} -> loc -| SexprInt {loc} -> loc -| SexprRec (loc, _, _) -> Some loc - -let parse_scope = function -| SexprRec (_, {loc;v=Some id}, toks) -> - if Id.Map.mem id !scope_table then - Id.Map.find id !scope_table toks - else - CErrors.user_err ?loc (str "Unknown scope" ++ spc () ++ Names.Id.print id) -| SexprStr {v=str} -> - let v_unit = CAst.make @@ CTacCst (AbsKn (Tuple 0)) in - ScopeRule (Extend.Atoken (Tok.PIDENT (Some str)), (fun _ -> v_unit)) -| tok -> - let loc = loc_of_token tok in - CErrors.user_err ?loc (str "Invalid parsing token") - -let parse_token = function -| SexprStr {v=s} -> TacTerm s -| SexprRec (_, {v=na}, [tok]) -> - let na = match na with None -> Anonymous | Some id -> Name id in - let scope = parse_scope tok in - TacNonTerm (na, scope) -| tok -> - let loc = loc_of_token tok in - CErrors.user_err ?loc (str "Invalid parsing token") - -end - -let parse_scope = ParseToken.parse_scope - -type synext = { - synext_tok : sexpr list; - synext_exp : raw_tacexpr; - synext_lev : int option; - synext_loc : bool; -} - -type krule = -| KRule : - (raw_tacexpr, _, 'act, Loc.t -> raw_tacexpr) Extend.rule * - ((Loc.t -> (Name.t * raw_tacexpr) list -> raw_tacexpr) -> 'act) -> krule - -let rec get_rule (tok : scope_rule token list) : krule = match tok with -| [] -> KRule (Extend.Stop, fun k loc -> k loc []) -| TacNonTerm (na, ScopeRule (scope, inj)) :: tok -> - let KRule (rule, act) = get_rule tok in - let rule = Extend.Next (rule, scope) in - let act k e = act (fun loc acc -> k loc ((na, inj e) :: acc)) in - KRule (rule, act) -| TacTerm t :: tok -> - let KRule (rule, act) = get_rule tok in - let rule = Extend.Next (rule, Extend.Atoken (CLexer.terminal t)) in - let act k _ = act k in - KRule (rule, act) - -let perform_notation syn st = - let tok = List.rev_map ParseToken.parse_token syn.synext_tok in - let KRule (rule, act) = get_rule tok in - let mk loc args = - let map (na, e) = - ((CAst.make ?loc:e.loc @@ CPatVar na), e) - in - let bnd = List.map map args in - CAst.make ~loc @@ CTacLet (false, bnd, syn.synext_exp) - in - let rule = Extend.Rule (rule, act mk) in - let lev = match syn.synext_lev with - | None -> None - | Some lev -> Some (string_of_int lev) - in - let rule = (lev, None, [rule]) in - ([Pcoq.ExtendRule (Pltac.tac2expr, None, (None, [rule]))], st) - -let ltac2_notation = - Pcoq.create_grammar_command "ltac2-notation" perform_notation - -let cache_synext (_, syn) = - Pcoq.extend_grammar_command ltac2_notation syn - -let open_synext i (_, syn) = - if Int.equal i 1 then Pcoq.extend_grammar_command ltac2_notation syn - -let subst_synext (subst, syn) = - let e = Tac2intern.subst_rawexpr subst syn.synext_exp in - if e == syn.synext_exp then syn else { syn with synext_exp = e } - -let classify_synext o = - if o.synext_loc then Dispose else Substitute o - -let inTac2Notation : synext -> obj = - declare_object {(default_object "TAC2-NOTATION") with - cache_function = cache_synext; - open_function = open_synext; - subst_function = subst_synext; - classify_function = classify_synext} - -type abbreviation = { - abbr_body : raw_tacexpr; -} - -let perform_abbreviation visibility ((sp, kn), abbr) = - let () = Tac2env.push_ltac visibility sp (TacAlias kn) in - Tac2env.define_alias kn abbr.abbr_body - -let load_abbreviation i obj = perform_abbreviation (Until i) obj -let open_abbreviation i obj = perform_abbreviation (Exactly i) obj - -let cache_abbreviation ((sp, kn), abbr) = - let () = Tac2env.push_ltac (Until 1) sp (TacAlias kn) in - Tac2env.define_alias kn abbr.abbr_body - -let subst_abbreviation (subst, abbr) = - let body' = subst_rawexpr subst abbr.abbr_body in - if body' == abbr.abbr_body then abbr - else { abbr_body = body' } - -let classify_abbreviation o = Substitute o - -let inTac2Abbreviation : abbreviation -> obj = - declare_object {(default_object "TAC2-ABBREVIATION") with - cache_function = cache_abbreviation; - load_function = load_abbreviation; - open_function = open_abbreviation; - subst_function = subst_abbreviation; - classify_function = classify_abbreviation} - -let register_notation ?(local = false) tkn lev body = match tkn, lev with -| [SexprRec (_, {loc;v=Some id}, [])], None -> - (* Tactic abbreviation *) - let () = check_lowercase CAst.(make ?loc id) in - let body = Tac2intern.globalize Id.Set.empty body in - let abbr = { abbr_body = body } in - ignore (Lib.add_leaf id (inTac2Abbreviation abbr)) -| _ -> - (* Check that the tokens make sense *) - let entries = List.map ParseToken.parse_token tkn in - let fold accu tok = match tok with - | TacTerm _ -> accu - | TacNonTerm (Name id, _) -> Id.Set.add id accu - | TacNonTerm (Anonymous, _) -> accu - in - let ids = List.fold_left fold Id.Set.empty entries in - (* Globalize so that names are absolute *) - let body = Tac2intern.globalize ids body in - let lev = match lev with Some _ -> lev | None -> Some 5 in - let ext = { - synext_tok = tkn; - synext_exp = body; - synext_lev = lev; - synext_loc = local; - } in - Lib.add_anonymous_leaf (inTac2Notation ext) - -type redefinition = { - redef_kn : ltac_constant; - redef_body : glb_tacexpr; -} - -let perform_redefinition (_, redef) = - let kn = redef.redef_kn in - let data = Tac2env.interp_global kn in - let data = { data with Tac2env.gdata_expr = redef.redef_body } in - Tac2env.define_global kn data - -let subst_redefinition (subst, redef) = - let kn = Mod_subst.subst_kn subst redef.redef_kn in - let body = Tac2intern.subst_expr subst redef.redef_body in - if kn == redef.redef_kn && body == redef.redef_body then redef - else { redef_kn = kn; redef_body = body } - -let classify_redefinition o = Substitute o - -let inTac2Redefinition : redefinition -> obj = - declare_object {(default_object "TAC2-REDEFINITION") with - cache_function = perform_redefinition; - open_function = (fun _ -> perform_redefinition); - subst_function = subst_redefinition; - classify_function = classify_redefinition } - -let register_redefinition ?(local = false) qid e = - let kn = - try Tac2env.locate_ltac qid - with Not_found -> user_err ?loc:qid.CAst.loc (str "Unknown tactic " ++ pr_qualid qid) - in - let kn = match kn with - | TacConstant kn -> kn - | TacAlias _ -> - user_err ?loc:qid.CAst.loc (str "Cannot redefine syntactic abbreviations") - in - let data = Tac2env.interp_global kn in - let () = - if not (data.Tac2env.gdata_mutable) then - user_err ?loc:qid.CAst.loc (str "The tactic " ++ pr_qualid qid ++ str " is not declared as mutable") - in - let (e, t) = intern ~strict:true e in - let () = - if not (is_value e) then - user_err ?loc:qid.CAst.loc (str "Tactic definition must be a syntactical value") - in - let () = - if not (Tac2intern.check_subtype t data.Tac2env.gdata_type) then - let name = int_name () in - user_err ?loc:qid.CAst.loc (str "Type " ++ pr_glbtype name (snd t) ++ - str " is not a subtype of " ++ pr_glbtype name (snd data.Tac2env.gdata_type)) - in - let def = { - redef_kn = kn; - redef_body = e; - } in - Lib.add_anonymous_leaf (inTac2Redefinition def) - -let perform_eval ~pstate e = - let open Proofview.Notations in - let env = Global.env () in - let (e, ty) = Tac2intern.intern ~strict:false e in - let v = Tac2interp.interp Tac2interp.empty_environment e in - let selector, proof = - match pstate with - | None -> - let sigma = Evd.from_env env in - let name, poly = Id.of_string "ltac2", false in - Goal_select.SelectAll, Proof.start ~name ~poly sigma [] - | Some pstate -> - Goal_select.get_default_goal_selector (), - Proof_global.give_me_the_proof pstate - in - let v = match selector with - | Goal_select.SelectNth i -> Proofview.tclFOCUS i i v - | Goal_select.SelectList l -> Proofview.tclFOCUSLIST l v - | Goal_select.SelectId id -> Proofview.tclFOCUSID id v - | Goal_select.SelectAll -> v - | Goal_select.SelectAlreadyFocused -> assert false (* TODO **) - in - (* HACK: the API doesn't allow to return a value *) - let ans = ref None in - let tac = (v >>= fun r -> ans := Some r; Proofview.tclUNIT ()) in - let (proof, _) = Proof.run_tactic (Global.env ()) tac proof in - let sigma = Proof.in_proof proof (fun sigma -> sigma) in - let ans = match !ans with None -> assert false | Some r -> r in - let name = int_name () in - Feedback.msg_notice (str "- : " ++ pr_glbtype name (snd ty) - ++ spc () ++ str "=" ++ spc () ++ - Tac2print.pr_valexpr env sigma ans (snd ty)) - -(** Toplevel entries *) - -let register_struct ?local ~pstate str = match str with -| StrVal (mut, isrec, e) -> register_ltac ?local ~mut isrec e -| StrTyp (isrec, t) -> register_type ?local isrec t -| StrPrm (id, t, ml) -> register_primitive ?local id t ml -| StrSyn (tok, lev, e) -> register_notation ?local tok lev e -| StrMut (qid, e) -> register_redefinition ?local qid e -| StrRun e -> perform_eval ~pstate e - -(** Toplevel exception *) - -let _ = Goptions.declare_bool_option { - Goptions.optdepr = false; - Goptions.optname = "print Ltac2 backtrace"; - Goptions.optkey = ["Ltac2"; "Backtrace"]; - Goptions.optread = (fun () -> !Tac2interp.print_ltac2_backtrace); - Goptions.optwrite = (fun b -> Tac2interp.print_ltac2_backtrace := b); -} - -let backtrace : backtrace Exninfo.t = Exninfo.make () - -let pr_frame = function -| FrAnon e -> str "Call {" ++ pr_glbexpr e ++ str "}" -| FrLtac kn -> - str "Call " ++ Libnames.pr_qualid (Tac2env.shortest_qualid_of_ltac (TacConstant kn)) -| FrPrim ml -> - str "Prim <" ++ str ml.mltac_plugin ++ str ":" ++ str ml.mltac_tactic ++ str ">" -| FrExtn (tag, arg) -> - let obj = Tac2env.interp_ml_object tag in - str "Extn " ++ str (Tac2dyn.Arg.repr tag) ++ str ":" ++ spc () ++ - obj.Tac2env.ml_print (Global.env ()) arg - -let () = register_handler begin function -| Tac2interp.LtacError (kn, args) -> - let t_exn = KerName.make Tac2env.coq_prefix (Label.make "exn") in - let v = Tac2ffi.of_open (kn, args) in - let t = GTypRef (Other t_exn, []) in - let c = Tac2print.pr_valexpr (Global.env ()) Evd.empty v t in - hov 0 (str "Uncaught Ltac2 exception:" ++ spc () ++ hov 0 c) -| _ -> raise Unhandled -end - -let () = ExplainErr.register_additional_error_info begin fun (e, info) -> - if !Tac2interp.print_ltac2_backtrace then - let bt = Exninfo.get info backtrace in - let bt = match bt with - | Some bt -> bt - | None -> raise Exit - in - let bt = - str "Backtrace:" ++ fnl () ++ prlist_with_sep fnl pr_frame bt ++ fnl () - in - Some (Loc.tag @@ Some bt) - else raise Exit -end - -(** Printing *) - -let print_ltac qid = - if Tac2env.is_constructor qid then - let kn = - try Tac2env.locate_constructor qid - with Not_found -> user_err ?loc:qid.CAst.loc (str "Unknown constructor " ++ pr_qualid qid) - in - let _ = Tac2env.interp_constructor kn in - Feedback.msg_notice (hov 2 (str "Constructor" ++ spc () ++ str ":" ++ spc () ++ pr_qualid qid)) - else - let kn = - try Tac2env.locate_ltac qid - with Not_found -> user_err ?loc:qid.CAst.loc (str "Unknown tactic " ++ pr_qualid qid) - in - match kn with - | TacConstant kn -> - let data = Tac2env.interp_global kn in - let e = data.Tac2env.gdata_expr in - let (_, t) = data.Tac2env.gdata_type in - let name = int_name () in - Feedback.msg_notice ( - hov 0 ( - hov 2 (pr_qualid qid ++ spc () ++ str ":" ++ spc () ++ pr_glbtype name t) ++ fnl () ++ - hov 2 (pr_qualid qid ++ spc () ++ str ":=" ++ spc () ++ pr_glbexpr e) - ) - ) - | TacAlias kn -> - Feedback.msg_notice (str "Alias to ...") - -(** Calling tactics *) - -let solve ~pstate default tac = - let pstate, status = Proof_global.with_current_proof begin fun etac p -> - let with_end_tac = if default then Some etac else None in - let g = Goal_select.get_default_goal_selector () in - let (p, status) = Pfedit.solve g None tac ?with_end_tac p in - (* in case a strict subtree was completed, - go back to the top of the prooftree *) - let p = Proof.maximal_unfocus Vernacentries.command_focus p in - p, status - end pstate in - if not status then Feedback.feedback Feedback.AddedAxiom; - pstate - -let call ~pstate ~default e = - let loc = e.loc in - let (e, t) = intern ~strict:false e in - let () = check_unit ?loc t in - let tac = Tac2interp.interp Tac2interp.empty_environment e in - solve ~pstate default (Proofview.tclIGNORE tac) - -(** Primitive algebraic types than can't be defined Coq-side *) - -let register_prim_alg name params def = - let id = Id.of_string name in - let def = List.map (fun (cstr, tpe) -> (Id.of_string_soft cstr, tpe)) def in - let getn (const, nonconst) (c, args) = match args with - | [] -> (succ const, nonconst) - | _ :: _ -> (const, succ nonconst) - in - let nconst, nnonconst = List.fold_left getn (0, 0) def in - let alg = { - galg_constructors = def; - galg_nconst = nconst; - galg_nnonconst = nnonconst; - } in - let def = (params, GTydAlg alg) in - let def = { typdef_local = false; typdef_expr = def } in - ignore (Lib.add_leaf id (inTypDef def)) - -let coq_def n = KerName.make Tac2env.coq_prefix (Label.make n) - -let def_unit = { - typdef_local = false; - typdef_expr = 0, GTydDef (Some (GTypRef (Tuple 0, []))); -} - -let t_list = coq_def "list" - -let (f_register_constr_quotations, register_constr_quotations) = Hook.make () - -let cache_ltac2_init (_, ()) = - Hook.get f_register_constr_quotations () - -let load_ltac2_init _ (_, ()) = - Hook.get f_register_constr_quotations () - -let open_ltac2_init _ (_, ()) = - Goptions.set_string_option_value_gen ["Default"; "Proof"; "Mode"] "Ltac2" - -(** Dummy object that register global rules when Require is called *) -let inTac2Init : unit -> obj = - declare_object {(default_object "TAC2-INIT") with - cache_function = cache_ltac2_init; - load_function = load_ltac2_init; - open_function = open_ltac2_init; - } - -let _ = Mltop.declare_cache_obj begin fun () -> - ignore (Lib.add_leaf (Id.of_string "unit") (inTypDef def_unit)); - register_prim_alg "list" 1 [ - ("[]", []); - ("::", [GTypVar 0; GTypRef (Other t_list, [GTypVar 0])]); - ]; - Lib.add_anonymous_leaf (inTac2Init ()); -end "ltac2_plugin" diff --git a/src/tac2entries.mli b/src/tac2entries.mli deleted file mode 100644 index d493192bb3..0000000000 --- a/src/tac2entries.mli +++ /dev/null @@ -1,93 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* ?mut:bool -> rec_flag -> - (Names.lname * raw_tacexpr) list -> unit - -val register_type : ?local:bool -> rec_flag -> - (qualid * redef_flag * raw_quant_typedef) list -> unit - -val register_primitive : ?local:bool -> - Names.lident -> raw_typexpr -> ml_tactic_name -> unit - -val register_struct - : ?local:bool - -> pstate:Proof_global.t option - -> strexpr - -> unit - -val register_notation : ?local:bool -> sexpr list -> int option -> - raw_tacexpr -> unit - -(** {5 Notations} *) - -type scope_rule = -| ScopeRule : (raw_tacexpr, _, 'a) Extend.symbol * ('a -> raw_tacexpr) -> scope_rule - -type scope_interpretation = sexpr list -> scope_rule - -val register_scope : Id.t -> scope_interpretation -> unit -(** Create a new scope with the provided name *) - -val parse_scope : sexpr -> scope_rule -(** Use this to interpret the subscopes for interpretation functions *) - -(** {5 Inspecting} *) - -val print_ltac : Libnames.qualid -> unit - -(** {5 Eval loop} *) - -(** Evaluate a tactic expression in the current environment *) -val call : pstate:Proof_global.t -> default:bool -> raw_tacexpr -> Proof_global.t - -(** {5 Toplevel exceptions} *) - -val backtrace : backtrace Exninfo.t - -(** {5 Parsing entries} *) - -module Pltac : -sig -val tac2expr : raw_tacexpr Pcoq.Entry.t - -(** Quoted entries. To be used for complex notations. *) - -open Tac2qexpr - -val q_ident : Id.t CAst.t or_anti Pcoq.Entry.t -val q_bindings : bindings Pcoq.Entry.t -val q_with_bindings : bindings Pcoq.Entry.t -val q_intropattern : intro_pattern Pcoq.Entry.t -val q_intropatterns : intro_pattern list CAst.t Pcoq.Entry.t -val q_destruction_arg : destruction_arg Pcoq.Entry.t -val q_induction_clause : induction_clause Pcoq.Entry.t -val q_conversion : conversion Pcoq.Entry.t -val q_rewriting : rewriting Pcoq.Entry.t -val q_clause : clause Pcoq.Entry.t -val q_dispatch : dispatch Pcoq.Entry.t -val q_occurrences : occurrences Pcoq.Entry.t -val q_reference : reference or_anti Pcoq.Entry.t -val q_strategy_flag : strategy_flag Pcoq.Entry.t -val q_constr_matching : constr_matching Pcoq.Entry.t -val q_goal_matching : goal_matching Pcoq.Entry.t -val q_hintdb : hintdb Pcoq.Entry.t -val q_move_location : move_location Pcoq.Entry.t -val q_pose : pose Pcoq.Entry.t -val q_assert : assertion Pcoq.Entry.t -end - -(** {5 Hooks} *) - -val register_constr_quotations : (unit -> unit) Hook.t diff --git a/src/tac2env.ml b/src/tac2env.ml deleted file mode 100644 index 93ad57e97e..0000000000 --- a/src/tac2env.ml +++ /dev/null @@ -1,298 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* KerName.compare c1 c2 -| TacAlias c1, TacAlias c2 -> KerName.compare c1 c2 -| TacConstant _, TacAlias _ -> -1 -| TacAlias _, TacConstant _ -> 1 - -let equal r1 r2 = compare r1 r2 == 0 - -end - -module KnTab = Nametab.Make(FullPath)(KerName) -module RfTab = Nametab.Make(FullPath)(TacRef) -module RfMap = Map.Make(TacRef) - -type nametab = { - tab_ltac : RfTab.t; - tab_ltac_rev : full_path RfMap.t; - tab_cstr : KnTab.t; - tab_cstr_rev : full_path KNmap.t; - tab_type : KnTab.t; - tab_type_rev : full_path KNmap.t; - tab_proj : KnTab.t; - tab_proj_rev : full_path KNmap.t; -} - -let empty_nametab = { - tab_ltac = RfTab.empty; - tab_ltac_rev = RfMap.empty; - tab_cstr = KnTab.empty; - tab_cstr_rev = KNmap.empty; - tab_type = KnTab.empty; - tab_type_rev = KNmap.empty; - tab_proj = KnTab.empty; - tab_proj_rev = KNmap.empty; -} - -let nametab = Summary.ref empty_nametab ~name:"ltac2-nametab" - -let push_ltac vis sp kn = - let tab = !nametab in - let tab_ltac = RfTab.push vis sp kn tab.tab_ltac in - let tab_ltac_rev = RfMap.add kn sp tab.tab_ltac_rev in - nametab := { tab with tab_ltac; tab_ltac_rev } - -let locate_ltac qid = - let tab = !nametab in - RfTab.locate qid tab.tab_ltac - -let locate_extended_all_ltac qid = - let tab = !nametab in - RfTab.find_prefixes qid tab.tab_ltac - -let shortest_qualid_of_ltac kn = - let tab = !nametab in - let sp = RfMap.find kn tab.tab_ltac_rev in - RfTab.shortest_qualid Id.Set.empty sp tab.tab_ltac - -let push_constructor vis sp kn = - let tab = !nametab in - let tab_cstr = KnTab.push vis sp kn tab.tab_cstr in - let tab_cstr_rev = KNmap.add kn sp tab.tab_cstr_rev in - nametab := { tab with tab_cstr; tab_cstr_rev } - -let locate_constructor qid = - let tab = !nametab in - KnTab.locate qid tab.tab_cstr - -let locate_extended_all_constructor qid = - let tab = !nametab in - KnTab.find_prefixes qid tab.tab_cstr - -let shortest_qualid_of_constructor kn = - let tab = !nametab in - let sp = KNmap.find kn tab.tab_cstr_rev in - KnTab.shortest_qualid Id.Set.empty sp tab.tab_cstr - -let push_type vis sp kn = - let tab = !nametab in - let tab_type = KnTab.push vis sp kn tab.tab_type in - let tab_type_rev = KNmap.add kn sp tab.tab_type_rev in - nametab := { tab with tab_type; tab_type_rev } - -let locate_type qid = - let tab = !nametab in - KnTab.locate qid tab.tab_type - -let locate_extended_all_type qid = - let tab = !nametab in - KnTab.find_prefixes qid tab.tab_type - -let shortest_qualid_of_type ?loc kn = - let tab = !nametab in - let sp = KNmap.find kn tab.tab_type_rev in - KnTab.shortest_qualid ?loc Id.Set.empty sp tab.tab_type - -let push_projection vis sp kn = - let tab = !nametab in - let tab_proj = KnTab.push vis sp kn tab.tab_proj in - let tab_proj_rev = KNmap.add kn sp tab.tab_proj_rev in - nametab := { tab with tab_proj; tab_proj_rev } - -let locate_projection qid = - let tab = !nametab in - KnTab.locate qid tab.tab_proj - -let locate_extended_all_projection qid = - let tab = !nametab in - KnTab.find_prefixes qid tab.tab_proj - -let shortest_qualid_of_projection kn = - let tab = !nametab in - let sp = KNmap.find kn tab.tab_proj_rev in - KnTab.shortest_qualid Id.Set.empty sp tab.tab_proj - -type 'a or_glb_tacexpr = -| GlbVal of 'a -| GlbTacexpr of glb_tacexpr - -type environment = { - env_ist : valexpr Id.Map.t; -} - -type ('a, 'b, 'r) intern_fun = Genintern.glob_sign -> 'a -> 'b * 'r glb_typexpr - -type ('a, 'b) ml_object = { - ml_intern : 'r. (raw_tacexpr, glb_tacexpr, 'r) intern_fun -> ('a, 'b or_glb_tacexpr, 'r) intern_fun; - ml_subst : Mod_subst.substitution -> 'b -> 'b; - ml_interp : environment -> 'b -> valexpr Proofview.tactic; - ml_print : Environ.env -> 'b -> Pp.t; -} - -module MLTypeObj = -struct - type ('a, 'b) t = ('a, 'b) ml_object -end - -module MLType = Tac2dyn.ArgMap(MLTypeObj) - -let ml_object_table = ref MLType.empty - -let define_ml_object t tpe = - ml_object_table := MLType.add t (MLType.Pack tpe) !ml_object_table - -let interp_ml_object t = - try - let MLType.Pack ans = MLType.find t !ml_object_table in - ans - with Not_found -> - CErrors.anomaly Pp.(str "Unknown object type " ++ str (Tac2dyn.Arg.repr t)) - -(** Absolute paths *) - -let coq_prefix = - MPfile (DirPath.make (List.map Id.of_string ["Init"; "Ltac2"])) - -let std_prefix = - MPfile (DirPath.make (List.map Id.of_string ["Std"; "Ltac2"])) - -let ltac1_prefix = - MPfile (DirPath.make (List.map Id.of_string ["Ltac1"; "Ltac2"])) - -(** Generic arguments *) - -let wit_ltac2 = Genarg.make0 "ltac2:value" -let wit_ltac2_quotation = Genarg.make0 "ltac2:quotation" -let () = Geninterp.register_val0 wit_ltac2 None -let () = Geninterp.register_val0 wit_ltac2_quotation None - -let is_constructor qid = - let (_, id) = repr_qualid qid in - let id = Id.to_string id in - assert (String.length id > 0); - match id with - | "true" | "false" -> true (* built-in constructors *) - | _ -> - match id.[0] with - | 'A'..'Z' -> true - | _ -> false diff --git a/src/tac2env.mli b/src/tac2env.mli deleted file mode 100644 index c7e87c5432..0000000000 --- a/src/tac2env.mli +++ /dev/null @@ -1,146 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* global_data -> unit -val interp_global : ltac_constant -> global_data - -(** {5 Toplevel definition of types} *) - -val define_type : type_constant -> glb_quant_typedef -> unit -val interp_type : type_constant -> glb_quant_typedef - -(** {5 Toplevel definition of algebraic constructors} *) - -type constructor_data = { - cdata_prms : int; - (** Type parameters *) - cdata_type : type_constant; - (** Inductive definition to which the constructor pertains *) - cdata_args : int glb_typexpr list; - (** Types of the constructor arguments *) - cdata_indx : int option; - (** Index of the constructor in the ADT. Numbering is duplicated between - argumentless and argument-using constructors, e.g. in type ['a option] - [None] and [Some] have both index 0. This field is empty whenever the - constructor is a member of an open type. *) -} - -val define_constructor : ltac_constructor -> constructor_data -> unit -val interp_constructor : ltac_constructor -> constructor_data - -(** {5 Toplevel definition of projections} *) - -type projection_data = { - pdata_prms : int; - (** Type parameters *) - pdata_type : type_constant; - (** Record definition to which the projection pertains *) - pdata_ptyp : int glb_typexpr; - (** Type of the projection *) - pdata_mutb : bool; - (** Whether the field is mutable *) - pdata_indx : int; - (** Index of the projection *) -} - -val define_projection : ltac_projection -> projection_data -> unit -val interp_projection : ltac_projection -> projection_data - -(** {5 Toplevel definition of aliases} *) - -val define_alias : ltac_constant -> raw_tacexpr -> unit -val interp_alias : ltac_constant -> raw_tacexpr - -(** {5 Name management} *) - -val push_ltac : visibility -> full_path -> tacref -> unit -val locate_ltac : qualid -> tacref -val locate_extended_all_ltac : qualid -> tacref list -val shortest_qualid_of_ltac : tacref -> qualid - -val push_constructor : visibility -> full_path -> ltac_constructor -> unit -val locate_constructor : qualid -> ltac_constructor -val locate_extended_all_constructor : qualid -> ltac_constructor list -val shortest_qualid_of_constructor : ltac_constructor -> qualid - -val push_type : visibility -> full_path -> type_constant -> unit -val locate_type : qualid -> type_constant -val locate_extended_all_type : qualid -> type_constant list -val shortest_qualid_of_type : ?loc:Loc.t -> type_constant -> qualid - -val push_projection : visibility -> full_path -> ltac_projection -> unit -val locate_projection : qualid -> ltac_projection -val locate_extended_all_projection : qualid -> ltac_projection list -val shortest_qualid_of_projection : ltac_projection -> qualid - -(** {5 Toplevel definitions of ML tactics} *) - -(** This state is not part of the summary, contrarily to the ones above. It is - intended to be used from ML plugins to register ML-side functions. *) - -val define_primitive : ml_tactic_name -> closure -> unit -val interp_primitive : ml_tactic_name -> closure - -(** {5 ML primitive types} *) - -type 'a or_glb_tacexpr = -| GlbVal of 'a -| GlbTacexpr of glb_tacexpr - -type ('a, 'b, 'r) intern_fun = Genintern.glob_sign -> 'a -> 'b * 'r glb_typexpr - -type environment = { - env_ist : valexpr Id.Map.t; -} - -type ('a, 'b) ml_object = { - ml_intern : 'r. (raw_tacexpr, glb_tacexpr, 'r) intern_fun -> ('a, 'b or_glb_tacexpr, 'r) intern_fun; - ml_subst : Mod_subst.substitution -> 'b -> 'b; - ml_interp : environment -> 'b -> valexpr Proofview.tactic; - ml_print : Environ.env -> 'b -> Pp.t; -} - -val define_ml_object : ('a, 'b) Tac2dyn.Arg.tag -> ('a, 'b) ml_object -> unit -val interp_ml_object : ('a, 'b) Tac2dyn.Arg.tag -> ('a, 'b) ml_object - -(** {5 Absolute paths} *) - -val coq_prefix : ModPath.t -(** Path where primitive datatypes are defined in Ltac2 plugin. *) - -val std_prefix : ModPath.t -(** Path where Ltac-specific datatypes are defined in Ltac2 plugin. *) - -val ltac1_prefix : ModPath.t -(** Path where the Ltac1 legacy FFI is defined. *) - -(** {5 Generic arguments} *) - -val wit_ltac2 : (raw_tacexpr, glb_tacexpr, Util.Empty.t) genarg_type -val wit_ltac2_quotation : (Id.t Loc.located, Id.t, Util.Empty.t) genarg_type - -(** {5 Helper functions} *) - -val is_constructor : qualid -> bool diff --git a/src/tac2expr.mli b/src/tac2expr.mli deleted file mode 100644 index 1069d0bfa3..0000000000 --- a/src/tac2expr.mli +++ /dev/null @@ -1,190 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* raw_tacexpr_r - -and raw_tacexpr = raw_tacexpr_r CAst.t - -and raw_taccase = raw_patexpr * raw_tacexpr - -and raw_recexpr = (ltac_projection or_relid * raw_tacexpr) list - -type case_info = type_constant or_tuple - -type 'a open_match = { - opn_match : 'a; - opn_branch : (Name.t * Name.t array * 'a) KNmap.t; - (** Invariant: should not be empty *) - opn_default : Name.t * 'a; -} - -type glb_tacexpr = -| GTacAtm of atom -| GTacVar of Id.t -| GTacRef of ltac_constant -| GTacFun of Name.t list * glb_tacexpr -| GTacApp of glb_tacexpr * glb_tacexpr list -| GTacLet of rec_flag * (Name.t * glb_tacexpr) list * glb_tacexpr -| GTacCst of case_info * int * glb_tacexpr list -| GTacCse of glb_tacexpr * case_info * glb_tacexpr array * (Name.t array * glb_tacexpr) array -| GTacPrj of type_constant * glb_tacexpr * int -| GTacSet of type_constant * glb_tacexpr * int * glb_tacexpr -| GTacOpn of ltac_constructor * glb_tacexpr list -| GTacWth of glb_tacexpr open_match -| GTacExt : (_, 'a) Tac2dyn.Arg.tag * 'a -> glb_tacexpr -| GTacPrm of ml_tactic_name * glb_tacexpr list - -(** {5 Parsing & Printing} *) - -type exp_level = -| E5 -| E4 -| E3 -| E2 -| E1 -| E0 - -type sexpr = -| SexprStr of string CAst.t -| SexprInt of int CAst.t -| SexprRec of Loc.t * Id.t option CAst.t * sexpr list - -(** {5 Toplevel statements} *) - -type strexpr = -| StrVal of mutable_flag * rec_flag * (Names.lname * raw_tacexpr) list - (** Term definition *) -| StrTyp of rec_flag * (qualid * redef_flag * raw_quant_typedef) list - (** Type definition *) -| StrPrm of Names.lident * raw_typexpr * ml_tactic_name - (** External definition *) -| StrSyn of sexpr list * int option * raw_tacexpr - (** Syntactic extensions *) -| StrMut of qualid * raw_tacexpr - (** Redefinition of mutable globals *) -| StrRun of raw_tacexpr - (** Toplevel evaluation of an expression *) - -(** {5 Dynamic semantics} *) - -(** Values are represented in a way similar to OCaml, i.e. they constrast - immediate integers (integers, constructors without arguments) and structured - blocks (tuples, arrays, constructors with arguments), as well as a few other - base cases, namely closures, strings, named constructors, and dynamic type - coming from the Coq implementation. *) - -type tag = int - -type frame = -| FrLtac of ltac_constant -| FrAnon of glb_tacexpr -| FrPrim of ml_tactic_name -| FrExtn : ('a, 'b) Tac2dyn.Arg.tag * 'b -> frame - -type backtrace = frame list diff --git a/src/tac2extffi.ml b/src/tac2extffi.ml deleted file mode 100644 index 315c970f9e..0000000000 --- a/src/tac2extffi.ml +++ /dev/null @@ -1,40 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* assert false) f - -(** More ML representations *) - -let to_qhyp v = match Value.to_block v with -| (0, [| i |]) -> AnonHyp (Value.to_int i) -| (1, [| id |]) -> NamedHyp (Value.to_ident id) -| _ -> assert false - -let qhyp = make_to_repr to_qhyp - -let to_bindings = function -| ValInt 0 -> NoBindings -| ValBlk (0, [| vl |]) -> - ImplicitBindings (Value.to_list Value.to_constr vl) -| ValBlk (1, [| vl |]) -> - ExplicitBindings ((Value.to_list (fun p -> to_pair to_qhyp Value.to_constr p) vl)) -| _ -> assert false - -let bindings = make_to_repr to_bindings - -let to_constr_with_bindings v = match Value.to_tuple v with -| [| c; bnd |] -> (Value.to_constr c, to_bindings bnd) -| _ -> assert false - -let constr_with_bindings = make_to_repr to_constr_with_bindings diff --git a/src/tac2extffi.mli b/src/tac2extffi.mli deleted file mode 100644 index f5251c3d0d..0000000000 --- a/src/tac2extffi.mli +++ /dev/null @@ -1,16 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* 'a Proofview.tactic) arity0 -| AddAty : ('a, 'b) arity0 -> ('a, 'a -> 'b) arity0 - -type valexpr = -| ValInt of int - (** Immediate integers *) -| ValBlk of tag * valexpr array - (** Structured blocks *) -| ValStr of Bytes.t - (** Strings *) -| ValCls of closure - (** Closures *) -| ValOpn of KerName.t * valexpr array - (** Open constructors *) -| ValExt : 'a Tac2dyn.Val.tag * 'a -> valexpr - (** Arbitrary data *) - -and closure = MLTactic : (valexpr, 'v) arity0 * 'v -> closure - -let arity_one = OneAty -let arity_suc a = AddAty a - -type 'a arity = (valexpr, 'a) arity0 - -let mk_closure arity f = MLTactic (arity, f) - -module Valexpr = -struct - -type t = valexpr - -let is_int = function -| ValInt _ -> true -| ValBlk _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ -> false - -let tag v = match v with -| ValBlk (n, _) -> n -| ValInt _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ -> - CErrors.anomaly (Pp.str "Unexpected value shape") - -let field v n = match v with -| ValBlk (_, v) -> v.(n) -| ValInt _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ -> - CErrors.anomaly (Pp.str "Unexpected value shape") - -let set_field v n w = match v with -| ValBlk (_, v) -> v.(n) <- w -| ValInt _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ -> - CErrors.anomaly (Pp.str "Unexpected value shape") - -let make_block tag v = ValBlk (tag, v) -let make_int n = ValInt n - -end - -type 'a repr = { - r_of : 'a -> valexpr; - r_to : valexpr -> 'a; - r_id : bool; -} - -let repr_of r x = r.r_of x -let repr_to r x = r.r_to x - -let make_repr r_of r_to = { r_of; r_to; r_id = false; } - -(** Dynamic tags *) - -let val_exn = Val.create "exn" -let val_constr = Val.create "constr" -let val_ident = Val.create "ident" -let val_pattern = Val.create "pattern" -let val_pp = Val.create "pp" -let val_sort = Val.create "sort" -let val_cast = Val.create "cast" -let val_inductive = Val.create "inductive" -let val_constant = Val.create "constant" -let val_constructor = Val.create "constructor" -let val_projection = Val.create "projection" -let val_case = Val.create "case" -let val_univ = Val.create "universe" -let val_free : Names.Id.Set.t Val.tag = Val.create "free" -let val_ltac1 : Geninterp.Val.t Val.tag = Val.create "ltac1" - -let extract_val (type a) (type b) (tag : a Val.tag) (tag' : b Val.tag) (v : b) : a = -match Val.eq tag tag' with -| None -> assert false -| Some Refl -> v - -(** Exception *) - -exception LtacError of KerName.t * valexpr array - -(** Conversion functions *) - -let valexpr = { - r_of = (fun obj -> obj); - r_to = (fun obj -> obj); - r_id = true; -} - -let of_unit () = ValInt 0 - -let to_unit = function -| ValInt 0 -> () -| _ -> assert false - -let unit = { - r_of = of_unit; - r_to = to_unit; - r_id = false; -} - -let of_int n = ValInt n -let to_int = function -| ValInt n -> n -| _ -> assert false - -let int = { - r_of = of_int; - r_to = to_int; - r_id = false; -} - -let of_bool b = if b then ValInt 0 else ValInt 1 - -let to_bool = function -| ValInt 0 -> true -| ValInt 1 -> false -| _ -> assert false - -let bool = { - r_of = of_bool; - r_to = to_bool; - r_id = false; -} - -let of_char n = ValInt (Char.code n) -let to_char = function -| ValInt n -> Char.chr n -| _ -> assert false - -let char = { - r_of = of_char; - r_to = to_char; - r_id = false; -} - -let of_string s = ValStr s -let to_string = function -| ValStr s -> s -| _ -> assert false - -let string = { - r_of = of_string; - r_to = to_string; - r_id = false; -} - -let rec of_list f = function -| [] -> ValInt 0 -| x :: l -> ValBlk (0, [| f x; of_list f l |]) - -let rec to_list f = function -| ValInt 0 -> [] -| ValBlk (0, [|v; vl|]) -> f v :: to_list f vl -| _ -> assert false - -let list r = { - r_of = (fun l -> of_list r.r_of l); - r_to = (fun l -> to_list r.r_to l); - r_id = false; -} - -let of_closure cls = ValCls cls - -let to_closure = function -| ValCls cls -> cls -| ValExt _ | ValInt _ | ValBlk _ | ValStr _ | ValOpn _ -> assert false - -let closure = { - r_of = of_closure; - r_to = to_closure; - r_id = false; -} - -let of_ext tag c = - ValExt (tag, c) - -let to_ext tag = function -| ValExt (tag', e) -> extract_val tag tag' e -| _ -> assert false - -let repr_ext tag = { - r_of = (fun e -> of_ext tag e); - r_to = (fun e -> to_ext tag e); - r_id = false; -} - -let of_constr c = of_ext val_constr c -let to_constr c = to_ext val_constr c -let constr = repr_ext val_constr - -let of_ident c = of_ext val_ident c -let to_ident c = to_ext val_ident c -let ident = repr_ext val_ident - -let of_pattern c = of_ext val_pattern c -let to_pattern c = to_ext val_pattern c -let pattern = repr_ext val_pattern - -let internal_err = - let open Names in - let coq_prefix = - MPfile (DirPath.make (List.map Id.of_string ["Init"; "Ltac2"])) - in - KerName.make coq_prefix (Label.of_id (Id.of_string "Internal")) - -(** FIXME: handle backtrace in Ltac2 exceptions *) -let of_exn c = match fst c with -| LtacError (kn, c) -> ValOpn (kn, c) -| _ -> ValOpn (internal_err, [|of_ext val_exn c|]) - -let to_exn c = match c with -| ValOpn (kn, c) -> - if Names.KerName.equal kn internal_err then - to_ext val_exn c.(0) - else - (LtacError (kn, c), Exninfo.null) -| _ -> assert false - -let exn = { - r_of = of_exn; - r_to = to_exn; - r_id = false; -} - -let of_option f = function -| None -> ValInt 0 -| Some c -> ValBlk (0, [|f c|]) - -let to_option f = function -| ValInt 0 -> None -| ValBlk (0, [|c|]) -> Some (f c) -| _ -> assert false - -let option r = { - r_of = (fun l -> of_option r.r_of l); - r_to = (fun l -> to_option r.r_to l); - r_id = false; -} - -let of_pp c = of_ext val_pp c -let to_pp c = to_ext val_pp c -let pp = repr_ext val_pp - -let of_tuple cl = ValBlk (0, cl) -let to_tuple = function -| ValBlk (0, cl) -> cl -| _ -> assert false - -let of_pair f g (x, y) = ValBlk (0, [|f x; g y|]) -let to_pair f g = function -| ValBlk (0, [|x; y|]) -> (f x, g y) -| _ -> assert false -let pair r0 r1 = { - r_of = (fun p -> of_pair r0.r_of r1.r_of p); - r_to = (fun p -> to_pair r0.r_to r1.r_to p); - r_id = false; -} - -let of_array f vl = ValBlk (0, Array.map f vl) -let to_array f = function -| ValBlk (0, vl) -> Array.map f vl -| _ -> assert false -let array r = { - r_of = (fun l -> of_array r.r_of l); - r_to = (fun l -> to_array r.r_to l); - r_id = false; -} - -let of_block (n, args) = ValBlk (n, args) -let to_block = function -| ValBlk (n, args) -> (n, args) -| _ -> assert false - -let block = { - r_of = of_block; - r_to = to_block; - r_id = false; -} - -let of_open (kn, args) = ValOpn (kn, args) - -let to_open = function -| ValOpn (kn, args) -> (kn, args) -| _ -> assert false - -let open_ = { - r_of = of_open; - r_to = to_open; - r_id = false; -} - -let of_constant c = of_ext val_constant c -let to_constant c = to_ext val_constant c -let constant = repr_ext val_constant - -let of_reference = function -| VarRef id -> ValBlk (0, [| of_ident id |]) -| ConstRef cst -> ValBlk (1, [| of_constant cst |]) -| IndRef ind -> ValBlk (2, [| of_ext val_inductive ind |]) -| ConstructRef cstr -> ValBlk (3, [| of_ext val_constructor cstr |]) - -let to_reference = function -| ValBlk (0, [| id |]) -> VarRef (to_ident id) -| ValBlk (1, [| cst |]) -> ConstRef (to_constant cst) -| ValBlk (2, [| ind |]) -> IndRef (to_ext val_inductive ind) -| ValBlk (3, [| cstr |]) -> ConstructRef (to_ext val_constructor cstr) -| _ -> assert false - -let reference = { - r_of = of_reference; - r_to = to_reference; - r_id = false; -} - -type ('a, 'b) fun1 = closure - -let fun1 (r0 : 'a repr) (r1 : 'b repr) : ('a, 'b) fun1 repr = closure -let to_fun1 r0 r1 f = to_closure f - -let rec apply : type a. a arity -> a -> valexpr list -> valexpr Proofview.tactic = - fun arity f args -> match args, arity with - | [], arity -> Proofview.tclUNIT (ValCls (MLTactic (arity, f))) - (* A few hardcoded cases for efficiency *) - | [a0], OneAty -> f a0 - | [a0; a1], AddAty OneAty -> f a0 a1 - | [a0; a1; a2], AddAty (AddAty OneAty) -> f a0 a1 a2 - | [a0; a1; a2; a3], AddAty (AddAty (AddAty OneAty)) -> f a0 a1 a2 a3 - (* Generic cases *) - | a :: args, OneAty -> - f a >>= fun f -> - let MLTactic (arity, f) = to_closure f in - apply arity f args - | a :: args, AddAty arity -> - apply arity (f a) args - -let apply (MLTactic (arity, f)) args = apply arity f args - -type n_closure = -| NClosure : 'a arity * (valexpr list -> 'a) -> n_closure - -let rec abstract n f = - if Int.equal n 1 then NClosure (OneAty, fun accu v -> f (List.rev (v :: accu))) - else - let NClosure (arity, fe) = abstract (n - 1) f in - NClosure (AddAty arity, fun accu v -> fe (v :: accu)) - -let abstract n f = - let () = assert (n > 0) in - let NClosure (arity, f) = abstract n f in - MLTactic (arity, f []) - -let app_fun1 cls r0 r1 x = - apply cls [r0.r_of x] >>= fun v -> Proofview.tclUNIT (r1.r_to v) diff --git a/src/tac2ffi.mli b/src/tac2ffi.mli deleted file mode 100644 index bfc93d99e6..0000000000 --- a/src/tac2ffi.mli +++ /dev/null @@ -1,189 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* valexpr - (** Arbitrary data *) - -type 'a arity - -val arity_one : (valexpr -> valexpr Proofview.tactic) arity -val arity_suc : 'a arity -> (valexpr -> 'a) arity - -val mk_closure : 'v arity -> 'v -> closure - -module Valexpr : -sig - type t = valexpr - val is_int : t -> bool - val tag : t -> int - val field : t -> int -> t - val set_field : t -> int -> t -> unit - val make_block : int -> t array -> t - val make_int : int -> t -end - -(** {5 Ltac2 FFI} *) - -type 'a repr - -val repr_of : 'a repr -> 'a -> valexpr -val repr_to : 'a repr -> valexpr -> 'a - -val make_repr : ('a -> valexpr) -> (valexpr -> 'a) -> 'a repr - -(** These functions allow to convert back and forth between OCaml and Ltac2 - data representation. The [to_*] functions raise an anomaly whenever the data - has not expected shape. *) - -val of_unit : unit -> valexpr -val to_unit : valexpr -> unit -val unit : unit repr - -val of_int : int -> valexpr -val to_int : valexpr -> int -val int : int repr - -val of_bool : bool -> valexpr -val to_bool : valexpr -> bool -val bool : bool repr - -val of_char : char -> valexpr -val to_char : valexpr -> char -val char : char repr - -val of_string : Bytes.t -> valexpr -val to_string : valexpr -> Bytes.t -val string : Bytes.t repr - -val of_list : ('a -> valexpr) -> 'a list -> valexpr -val to_list : (valexpr -> 'a) -> valexpr -> 'a list -val list : 'a repr -> 'a list repr - -val of_constr : EConstr.t -> valexpr -val to_constr : valexpr -> EConstr.t -val constr : EConstr.t repr - -val of_exn : Exninfo.iexn -> valexpr -val to_exn : valexpr -> Exninfo.iexn -val exn : Exninfo.iexn repr - -val of_ident : Id.t -> valexpr -val to_ident : valexpr -> Id.t -val ident : Id.t repr - -val of_closure : closure -> valexpr -val to_closure : valexpr -> closure -val closure : closure repr - -val of_block : (int * valexpr array) -> valexpr -val to_block : valexpr -> (int * valexpr array) -val block : (int * valexpr array) repr - -val of_array : ('a -> valexpr) -> 'a array -> valexpr -val to_array : (valexpr -> 'a) -> valexpr -> 'a array -val array : 'a repr -> 'a array repr - -val of_tuple : valexpr array -> valexpr -val to_tuple : valexpr -> valexpr array - -val of_pair : ('a -> valexpr) -> ('b -> valexpr) -> 'a * 'b -> valexpr -val to_pair : (valexpr -> 'a) -> (valexpr -> 'b) -> valexpr -> 'a * 'b -val pair : 'a repr -> 'b repr -> ('a * 'b) repr - -val of_option : ('a -> valexpr) -> 'a option -> valexpr -val to_option : (valexpr -> 'a) -> valexpr -> 'a option -val option : 'a repr -> 'a option repr - -val of_pattern : Pattern.constr_pattern -> valexpr -val to_pattern : valexpr -> Pattern.constr_pattern -val pattern : Pattern.constr_pattern repr - -val of_pp : Pp.t -> valexpr -val to_pp : valexpr -> Pp.t -val pp : Pp.t repr - -val of_constant : Constant.t -> valexpr -val to_constant : valexpr -> Constant.t -val constant : Constant.t repr - -val of_reference : GlobRef.t -> valexpr -val to_reference : valexpr -> GlobRef.t -val reference : GlobRef.t repr - -val of_ext : 'a Val.tag -> 'a -> valexpr -val to_ext : 'a Val.tag -> valexpr -> 'a -val repr_ext : 'a Val.tag -> 'a repr - -val of_open : KerName.t * valexpr array -> valexpr -val to_open : valexpr -> KerName.t * valexpr array -val open_ : (KerName.t * valexpr array) repr - -type ('a, 'b) fun1 - -val app_fun1 : ('a, 'b) fun1 -> 'a repr -> 'b repr -> 'a -> 'b Proofview.tactic - -val to_fun1 : 'a repr -> 'b repr -> valexpr -> ('a, 'b) fun1 -val fun1 : 'a repr -> 'b repr -> ('a, 'b) fun1 repr - -val valexpr : valexpr repr - -(** {5 Dynamic tags} *) - -val val_constr : EConstr.t Val.tag -val val_ident : Id.t Val.tag -val val_pattern : Pattern.constr_pattern Val.tag -val val_pp : Pp.t Val.tag -val val_sort : ESorts.t Val.tag -val val_cast : Constr.cast_kind Val.tag -val val_inductive : inductive Val.tag -val val_constant : Constant.t Val.tag -val val_constructor : constructor Val.tag -val val_projection : Projection.t Val.tag -val val_case : Constr.case_info Val.tag -val val_univ : Univ.Level.t Val.tag -val val_free : Id.Set.t Val.tag -val val_ltac1 : Geninterp.Val.t Val.tag - -val val_exn : Exninfo.iexn Tac2dyn.Val.tag -(** Toplevel representation of OCaml exceptions. Invariant: no [LtacError] - should be put into a value with tag [val_exn]. *) - -(** Closures *) - -val apply : closure -> valexpr list -> valexpr Proofview.tactic -(** Given a closure, apply it to some arguments. Handling of argument mismatches - is done automatically, i.e. in case of over or under-application. *) - -val abstract : int -> (valexpr list -> valexpr Proofview.tactic) -> closure -(** Turn a fixed-arity function into a closure. The inner function is guaranteed - to be applied to a list whose size is the integer argument. *) - -(** Exception *) - -exception LtacError of KerName.t * valexpr array -(** Ltac2-defined exceptions seen from OCaml side *) diff --git a/src/tac2intern.ml b/src/tac2intern.ml deleted file mode 100644 index de99fb167f..0000000000 --- a/src/tac2intern.ml +++ /dev/null @@ -1,1545 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* elt -> bool -val create : unit -> 'a t -val fresh : 'a t -> elt -val find : elt -> 'a t -> (elt * 'a option) -val union : elt -> elt -> 'a t -> unit -val set : elt -> 'a -> 'a t -> unit -module Map : -sig - type key = elt - type +'a t - val empty : 'a t - val add : key -> 'a -> 'a t -> 'a t - val mem : key -> 'a t -> bool - val find : key -> 'a t -> 'a - val exists : (key -> 'a -> bool) -> 'a t -> bool -end -end -= -struct -type elt = int -let equal = Int.equal -module Map = Int.Map - -type 'a node = -| Canon of int * 'a option -| Equiv of elt - -type 'a t = { - mutable uf_data : 'a node array; - mutable uf_size : int; -} - -let resize p = - if Int.equal (Array.length p.uf_data) p.uf_size then begin - let nsize = 2 * p.uf_size + 1 in - let v = Array.make nsize (Equiv 0) in - Array.blit p.uf_data 0 v 0 (Array.length p.uf_data); - p.uf_data <- v; - end - -let create () = { uf_data = [||]; uf_size = 0 } - -let fresh p = - resize p; - let n = p.uf_size in - p.uf_data.(n) <- (Canon (1, None)); - p.uf_size <- n + 1; - n - -let rec lookup n p = - let node = Array.get p.uf_data n in - match node with - | Canon (size, v) -> n, size, v - | Equiv y -> - let ((z, _, _) as res) = lookup y p in - if not (Int.equal z y) then Array.set p.uf_data n (Equiv z); - res - -let find n p = - let (x, _, v) = lookup n p in (x, v) - -let union x y p = - let ((x, size1, _) as xcan) = lookup x p in - let ((y, size2, _) as ycan) = lookup y p in - let xcan, ycan = if size1 < size2 then xcan, ycan else ycan, xcan in - let x, _, xnode = xcan in - let y, _, ynode = ycan in - assert (Option.is_empty xnode); - assert (Option.is_empty ynode); - p.uf_data.(x) <- Equiv y; - p.uf_data.(y) <- Canon (size1 + size2, None) - -let set x v p = - let (x, s, v') = lookup x p in - assert (Option.is_empty v'); - p.uf_data.(x) <- Canon (s, Some v) - -end - -type mix_var = -| GVar of UF.elt -| LVar of int - -type mix_type_scheme = int * mix_var glb_typexpr - -type environment = { - env_var : mix_type_scheme Id.Map.t; - (** Type schemes of bound variables *) - env_cst : UF.elt glb_typexpr UF.t; - (** Unification state *) - env_als : UF.elt Id.Map.t ref; - (** Map user-facing type variables to unification variables *) - env_opn : bool; - (** Accept unbound type variables *) - env_rec : (KerName.t * int) Id.Map.t; - (** Recursive type definitions *) - env_str : bool; - (** True iff in strict mode *) -} - -let empty_env () = { - env_var = Id.Map.empty; - env_cst = UF.create (); - env_als = ref Id.Map.empty; - env_opn = true; - env_rec = Id.Map.empty; - env_str = true; -} - -let env_name env = - (* Generate names according to a provided environment *) - let mk num = - let base = num mod 26 in - let rem = num / 26 in - let name = String.make 1 (Char.chr (97 + base)) in - let suff = if Int.equal rem 0 then "" else string_of_int rem in - let name = name ^ suff in - name - in - let fold id elt acc = UF.Map.add elt (Id.to_string id) acc in - let vars = Id.Map.fold fold env.env_als.contents UF.Map.empty in - let vars = ref vars in - let rec fresh n = - let name = mk n in - if UF.Map.exists (fun _ name' -> String.equal name name') !vars then fresh (succ n) - else name - in - fun n -> - if UF.Map.mem n !vars then UF.Map.find n !vars - else - let ans = fresh 0 in - let () = vars := UF.Map.add n ans !vars in - ans - -let ltac2_env : environment Genintern.Store.field = - Genintern.Store.field () - -let drop_ltac2_env store = - Genintern.Store.remove store ltac2_env - -let fresh_id env = UF.fresh env.env_cst - -let get_alias {loc;v=id} env = - try Id.Map.find id env.env_als.contents - with Not_found -> - if env.env_opn then - let n = fresh_id env in - let () = env.env_als := Id.Map.add id n env.env_als.contents in - n - else user_err ?loc (str "Unbound type parameter " ++ Id.print id) - -let push_name id t env = match id with -| Anonymous -> env -| Name id -> { env with env_var = Id.Map.add id t env.env_var } - -let error_nargs_mismatch ?loc kn nargs nfound = - let cstr = Tac2env.shortest_qualid_of_constructor kn in - user_err ?loc (str "Constructor " ++ pr_qualid cstr ++ str " expects " ++ - int nargs ++ str " arguments, but is applied to " ++ int nfound ++ - str " arguments") - -let error_nparams_mismatch ?loc nargs nfound = - user_err ?loc (str "Type expects " ++ int nargs ++ - str " arguments, but is applied to " ++ int nfound ++ - str " arguments") - -let rec subst_type subst (t : 'a glb_typexpr) = match t with -| GTypVar id -> subst id -| GTypArrow (t1, t2) -> GTypArrow (subst_type subst t1, subst_type subst t2) -| GTypRef (qid, args) -> - GTypRef (qid, List.map (fun t -> subst_type subst t) args) - -let rec intern_type env ({loc;v=t} : raw_typexpr) : UF.elt glb_typexpr = match t with -| CTypVar (Name id) -> GTypVar (get_alias (CAst.make ?loc id) env) -| CTypVar Anonymous -> GTypVar (fresh_id env) -| CTypRef (rel, args) -> - let (kn, nparams) = match rel with - | RelId qid -> - let id = qualid_basename qid in - if qualid_is_ident qid && Id.Map.mem id env.env_rec then - let (kn, n) = Id.Map.find id env.env_rec in - (Other kn, n) - else - let kn = - try Tac2env.locate_type qid - with Not_found -> - user_err ?loc (str "Unbound type constructor " ++ pr_qualid qid) - in - let (nparams, _) = Tac2env.interp_type kn in - (Other kn, nparams) - | AbsKn (Other kn) -> - let (nparams, _) = Tac2env.interp_type kn in - (Other kn, nparams) - | AbsKn (Tuple n) -> - (Tuple n, n) - in - let nargs = List.length args in - let () = - if not (Int.equal nparams nargs) then - let qid = match rel with - | RelId lid -> lid - | AbsKn (Other kn) -> shortest_qualid_of_type ?loc kn - | AbsKn (Tuple _) -> assert false - in - user_err ?loc (strbrk "The type constructor " ++ pr_qualid qid ++ - strbrk " expects " ++ int nparams ++ strbrk " argument(s), but is here \ - applied to " ++ int nargs ++ strbrk "argument(s)") - in - GTypRef (kn, List.map (fun t -> intern_type env t) args) -| CTypArrow (t1, t2) -> GTypArrow (intern_type env t1, intern_type env t2) - -let fresh_type_scheme env (t : type_scheme) : UF.elt glb_typexpr = - let (n, t) = t in - let subst = Array.init n (fun _ -> fresh_id env) in - let substf i = GTypVar subst.(i) in - subst_type substf t - -let fresh_mix_type_scheme env (t : mix_type_scheme) : UF.elt glb_typexpr = - let (n, t) = t in - let subst = Array.init n (fun _ -> fresh_id env) in - let substf = function - | LVar i -> GTypVar subst.(i) - | GVar n -> GTypVar n - in - subst_type substf t - -let fresh_reftype env (kn : KerName.t or_tuple) = - let n = match kn with - | Other kn -> fst (Tac2env.interp_type kn) - | Tuple n -> n - in - let subst = Array.init n (fun _ -> fresh_id env) in - let t = GTypRef (kn, Array.map_to_list (fun i -> GTypVar i) subst) in - (subst, t) - -(** First-order unification algorithm *) -let is_unfoldable kn = match snd (Tac2env.interp_type kn) with -| GTydDef (Some _) -> true -| GTydDef None | GTydAlg _ | GTydRec _ | GTydOpn -> false - -let unfold env kn args = - let (nparams, def) = Tac2env.interp_type kn in - let def = match def with - | GTydDef (Some t) -> t - | _ -> assert false - in - let args = Array.of_list args in - let subst n = args.(n) in - subst_type subst def - -(** View function, allows to ensure head normal forms *) -let rec kind env t = match t with -| GTypVar id -> - let (id, v) = UF.find id env.env_cst in - begin match v with - | None -> GTypVar id - | Some t -> kind env t - end -| GTypRef (Other kn, tl) -> - if is_unfoldable kn then kind env (unfold env kn tl) else t -| GTypArrow _ | GTypRef (Tuple _, _) -> t - -(** Normalize unification variables without unfolding type aliases *) -let rec nf env t = match t with -| GTypVar id -> - let (id, v) = UF.find id env.env_cst in - begin match v with - | None -> GTypVar id - | Some t -> nf env t - end -| GTypRef (kn, tl) -> - let tl = List.map (fun t -> nf env t) tl in - GTypRef (kn, tl) -| GTypArrow (t, u) -> - let t = nf env t in - let u = nf env u in - GTypArrow (t, u) - -let pr_glbtype env t = - let t = nf env t in - let name = env_name env in - pr_glbtype name t - -exception Occur - -let rec occur_check env id t = match kind env t with -| GTypVar id' -> if UF.equal id id' then raise Occur -| GTypArrow (t1, t2) -> - let () = occur_check env id t1 in - occur_check env id t2 -| GTypRef (kn, tl) -> - List.iter (fun t -> occur_check env id t) tl - -exception CannotUnify of UF.elt glb_typexpr * UF.elt glb_typexpr - -let unify_var env id t = match kind env t with -| GTypVar id' -> - if not (UF.equal id id') then UF.union id id' env.env_cst -| GTypArrow _ | GTypRef _ -> - try - let () = occur_check env id t in - UF.set id t env.env_cst - with Occur -> raise (CannotUnify (GTypVar id, t)) - -let eq_or_tuple eq t1 t2 = match t1, t2 with -| Tuple n1, Tuple n2 -> Int.equal n1 n2 -| Other o1, Other o2 -> eq o1 o2 -| _ -> false - -let rec unify0 env t1 t2 = match kind env t1, kind env t2 with -| GTypVar id, t | t, GTypVar id -> - unify_var env id t -| GTypArrow (t1, u1), GTypArrow (t2, u2) -> - let () = unify0 env t1 t2 in - unify0 env u1 u2 -| GTypRef (kn1, tl1), GTypRef (kn2, tl2) -> - if eq_or_tuple KerName.equal kn1 kn2 then - List.iter2 (fun t1 t2 -> unify0 env t1 t2) tl1 tl2 - else raise (CannotUnify (t1, t2)) -| _ -> raise (CannotUnify (t1, t2)) - -let unify ?loc env t1 t2 = - try unify0 env t1 t2 - with CannotUnify (u1, u2) -> - user_err ?loc (str "This expression has type" ++ spc () ++ pr_glbtype env t1 ++ - spc () ++ str "but an expression was expected of type" ++ spc () ++ pr_glbtype env t2) - -let unify_arrow ?loc env ft args = - let ft0 = ft in - let rec iter ft args is_fun = match kind env ft, args with - | t, [] -> t - | GTypArrow (t1, ft), (loc, t2) :: args -> - let () = unify ?loc env t2 t1 in - iter ft args true - | GTypVar id, (_, t) :: args -> - let ft = GTypVar (fresh_id env) in - let () = unify_var env id (GTypArrow (t, ft)) in - iter ft args true - | GTypRef _, _ :: _ -> - if is_fun then - user_err ?loc (str "This function has type" ++ spc () ++ pr_glbtype env ft0 ++ - spc () ++ str "and is applied to too many arguments") - else - user_err ?loc (str "This expression has type" ++ spc () ++ pr_glbtype env ft0 ++ - spc () ++ str "and is not a function") - in - iter ft args false - -(** Term typing *) - -let is_pure_constructor kn = - match snd (Tac2env.interp_type kn) with - | GTydAlg _ | GTydOpn -> true - | GTydRec fields -> - let is_pure (_, mut, _) = not mut in - List.for_all is_pure fields - | GTydDef _ -> assert false (** Type definitions have no constructors *) - -let rec is_value = function -| GTacAtm (AtmInt _) | GTacVar _ | GTacRef _ | GTacFun _ -> true -| GTacAtm (AtmStr _) | GTacApp _ | GTacLet _ -> false -| GTacCst (Tuple _, _, el) -> List.for_all is_value el -| GTacCst (_, _, []) -> true -| GTacOpn (_, el) -> List.for_all is_value el -| GTacCst (Other kn, _, el) -> is_pure_constructor kn && List.for_all is_value el -| GTacCse _ | GTacPrj _ | GTacSet _ | GTacExt _ | GTacPrm _ -| GTacWth _ -> false - -let is_rec_rhs = function -| GTacFun _ -> true -| GTacAtm _ | GTacVar _ | GTacRef _ | GTacApp _ | GTacLet _ | GTacPrj _ -| GTacSet _ | GTacExt _ | GTacPrm _ | GTacCst _ -| GTacCse _ | GTacOpn _ | GTacWth _ -> false - -let rec fv_type f t accu = match t with -| GTypVar id -> f id accu -| GTypArrow (t1, t2) -> fv_type f t1 (fv_type f t2 accu) -| GTypRef (kn, tl) -> List.fold_left (fun accu t -> fv_type f t accu) accu tl - -let fv_env env = - let rec f id accu = match UF.find id env.env_cst with - | id, None -> UF.Map.add id () accu - | _, Some t -> fv_type f t accu - in - let fold_var id (_, t) accu = - let fmix id accu = match id with - | LVar _ -> accu - | GVar id -> f id accu - in - fv_type fmix t accu - in - let fv_var = Id.Map.fold fold_var env.env_var UF.Map.empty in - let fold_als _ id accu = f id accu in - Id.Map.fold fold_als !(env.env_als) fv_var - -let abstract_var env (t : UF.elt glb_typexpr) : mix_type_scheme = - let fv = fv_env env in - let count = ref 0 in - let vars = ref UF.Map.empty in - let rec subst id = - let (id, t) = UF.find id env.env_cst in - match t with - | None -> - if UF.Map.mem id fv then GTypVar (GVar id) - else - begin try UF.Map.find id !vars - with Not_found -> - let n = !count in - let var = GTypVar (LVar n) in - let () = incr count in - let () = vars := UF.Map.add id var !vars in - var - end - | Some t -> subst_type subst t - in - let t = subst_type subst t in - (!count, t) - -let monomorphic (t : UF.elt glb_typexpr) : mix_type_scheme = - let subst id = GTypVar (GVar id) in - (0, subst_type subst t) - -let warn_not_unit = - CWarnings.create ~name:"not-unit" ~category:"ltac" - (fun () -> strbrk "The following expression should have type unit.") - -let warn_redundant_clause = - CWarnings.create ~name:"redundant-clause" ~category:"ltac" - (fun () -> strbrk "The following clause is redundant.") - -let check_elt_unit loc env t = - let maybe_unit = match kind env t with - | GTypVar _ -> true - | GTypArrow _ -> false - | GTypRef (Tuple 0, []) -> true - | GTypRef _ -> false - in - if not maybe_unit then warn_not_unit ?loc () - -let check_elt_empty loc env t = match kind env t with -| GTypVar _ -> - user_err ?loc (str "Cannot infer an empty type for this expression") -| GTypArrow _ | GTypRef (Tuple _, _) -> - user_err ?loc (str "Type" ++ spc () ++ pr_glbtype env t ++ spc () ++ str "is not an empty type") -| GTypRef (Other kn, _) -> - let def = Tac2env.interp_type kn in - match def with - | _, GTydAlg { galg_constructors = [] } -> kn - | _ -> - user_err ?loc (str "Type" ++ spc () ++ pr_glbtype env t ++ spc () ++ str "is not an empty type") - -let check_unit ?loc t = - let env = empty_env () in - (* Should not matter, t should be closed. *) - let t = fresh_type_scheme env t in - let maybe_unit = match kind env t with - | GTypVar _ -> true - | GTypArrow _ -> false - | GTypRef (Tuple 0, []) -> true - | GTypRef _ -> false - in - if not maybe_unit then warn_not_unit ?loc () - -let check_redundant_clause = function -| [] -> () -| (p, _) :: _ -> warn_redundant_clause ?loc:p.loc () - -let get_variable0 mem var = match var with -| RelId qid -> - let id = qualid_basename qid in - if qualid_is_ident qid && mem id then ArgVar CAst.(make ?loc:qid.CAst.loc id) - else - let kn = - try Tac2env.locate_ltac qid - with Not_found -> - CErrors.user_err ?loc:qid.CAst.loc (str "Unbound value " ++ pr_qualid qid) - in - ArgArg kn -| AbsKn kn -> ArgArg kn - -let get_variable env var = - let mem id = Id.Map.mem id env.env_var in - get_variable0 mem var - -let get_constructor env var = match var with -| RelId qid -> - let c = try Some (Tac2env.locate_constructor qid) with Not_found -> None in - begin match c with - | Some knc -> Other knc - | None -> - CErrors.user_err ?loc:qid.CAst.loc (str "Unbound constructor " ++ pr_qualid qid) - end -| AbsKn knc -> knc - -let get_projection var = match var with -| RelId qid -> - let kn = try Tac2env.locate_projection qid with Not_found -> - user_err ?loc:qid.CAst.loc (pr_qualid qid ++ str " is not a projection") - in - Tac2env.interp_projection kn -| AbsKn kn -> - Tac2env.interp_projection kn - -let intern_atm env = function -| AtmInt n -> (GTacAtm (AtmInt n), GTypRef (Other t_int, [])) -| AtmStr s -> (GTacAtm (AtmStr s), GTypRef (Other t_string, [])) - -let invalid_pattern ?loc kn kn' = - let pr t = match t with - | Other kn' -> str "type " ++ pr_typref kn' - | Tuple n -> str "tuple of size " ++ int n - in - user_err ?loc (str "Invalid pattern, expected a pattern for " ++ - pr kn ++ str ", found a pattern for " ++ pr kn') (** FIXME *) - -(** Pattern view *) - -type glb_patexpr = -| GPatVar of Name.t -| GPatRef of ltac_constructor or_tuple * glb_patexpr list - -let rec intern_patexpr env {loc;v=pat} = match pat with -| CPatVar na -> GPatVar na -| CPatRef (qid, pl) -> - let kn = get_constructor env qid in - GPatRef (kn, List.map (fun p -> intern_patexpr env p) pl) -| CPatCnv (pat, ty) -> - user_err ?loc (str "Pattern not handled yet") - -type pattern_kind = -| PKind_empty -| PKind_variant of type_constant or_tuple -| PKind_open of type_constant -| PKind_any - -let get_pattern_kind env pl = match pl with -| [] -> PKind_empty -| p :: pl -> - let rec get_kind (p, _) pl = match intern_patexpr env p with - | GPatVar _ -> - begin match pl with - | [] -> PKind_any - | p :: pl -> get_kind p pl - end - | GPatRef (Other kn, pl) -> - let data = Tac2env.interp_constructor kn in - if Option.is_empty data.cdata_indx then PKind_open data.cdata_type - else PKind_variant (Other data.cdata_type) - | GPatRef (Tuple _, tp) -> PKind_variant (Tuple (List.length tp)) - in - get_kind p pl - -(** Internalization *) - -(** Used to generate a fresh tactic variable for pattern-expansion *) -let fresh_var avoid = - let bad id = - Id.Set.mem id avoid || - (try ignore (locate_ltac (qualid_of_ident id)); true with Not_found -> false) - in - Namegen.next_ident_away_from (Id.of_string "p") bad - -let add_name accu = function -| Name id -> Id.Set.add id accu -| Anonymous -> accu - -let rec ids_of_pattern accu {v=pat} = match pat with -| CPatVar Anonymous -> accu -| CPatVar (Name id) -> Id.Set.add id accu -| CPatRef (_, pl) -> - List.fold_left ids_of_pattern accu pl -| CPatCnv (pat, _) -> ids_of_pattern accu pat - -let loc_of_relid = function -| RelId {loc} -> loc -| AbsKn _ -> None - -let extract_pattern_type ({loc;v=p} as pat) = match p with -| CPatCnv (pat, ty) -> pat, Some ty -| CPatVar _ | CPatRef _ -> pat, None - -(** Expand pattern: [p => t] becomes [x => match x with p => t end] *) -let expand_pattern avoid bnd = - let fold (avoid, bnd) (pat, t) = - let na, expand = match pat.v with - | CPatVar na -> - (* Don't expand variable patterns *) - na, None - | _ -> - let id = fresh_var avoid in - let qid = RelId (qualid_of_ident ?loc:pat.loc id) in - Name id, Some qid - in - let avoid = ids_of_pattern avoid pat in - let avoid = add_name avoid na in - (avoid, (na, pat, expand) :: bnd) - in - let (_, bnd) = List.fold_left fold (avoid, []) bnd in - let fold e (na, pat, expand) = match expand with - | None -> e - | Some qid -> - let loc = loc_of_relid qid in - CAst.make ?loc @@ CTacCse (CAst.make ?loc @@ CTacRef qid, [pat, e]) - in - let expand e = List.fold_left fold e bnd in - let nas = List.rev_map (fun (na, _, _) -> na) bnd in - (nas, expand) - -let is_alias env qid = match get_variable env qid with -| ArgArg (TacAlias _) -> true -| ArgVar _ | (ArgArg (TacConstant _)) -> false - -let rec intern_rec env {loc;v=e} = match e with -| CTacAtm atm -> intern_atm env atm -| CTacRef qid -> - begin match get_variable env qid with - | ArgVar {CAst.v=id} -> - let sch = Id.Map.find id env.env_var in - (GTacVar id, fresh_mix_type_scheme env sch) - | ArgArg (TacConstant kn) -> - let { Tac2env.gdata_type = sch } = - try Tac2env.interp_global kn - with Not_found -> - CErrors.anomaly (str "Missing hardwired primitive " ++ KerName.print kn) - in - (GTacRef kn, fresh_type_scheme env sch) - | ArgArg (TacAlias kn) -> - let e = - try Tac2env.interp_alias kn - with Not_found -> - CErrors.anomaly (str "Missing hardwired alias " ++ KerName.print kn) - in - intern_rec env e - end -| CTacCst qid -> - let kn = get_constructor env qid in - intern_constructor env loc kn [] -| CTacFun (bnd, e) -> - let bnd = List.map extract_pattern_type bnd in - let map (_, t) = match t with - | None -> GTypVar (fresh_id env) - | Some t -> intern_type env t - in - let tl = List.map map bnd in - let (nas, exp) = expand_pattern (Id.Map.domain env.env_var) bnd in - let env = List.fold_left2 (fun env na t -> push_name na (monomorphic t) env) env nas tl in - let (e, t) = intern_rec env (exp e) in - let t = List.fold_right (fun t accu -> GTypArrow (t, accu)) tl t in - (GTacFun (nas, e), t) -| CTacApp ({loc;v=CTacCst qid}, args) -> - let kn = get_constructor env qid in - intern_constructor env loc kn args -| CTacApp ({v=CTacRef qid}, args) when is_alias env qid -> - let kn = match get_variable env qid with - | ArgArg (TacAlias kn) -> kn - | ArgVar _ | (ArgArg (TacConstant _)) -> assert false - in - let e = Tac2env.interp_alias kn in - let map arg = - (* Thunk alias arguments *) - let loc = arg.loc in - let t_unit = CAst.make ?loc @@ CTypRef (AbsKn (Tuple 0), []) in - let var = CAst.make ?loc @@ CPatCnv (CAst.make ?loc @@ CPatVar Anonymous, t_unit) in - CAst.make ?loc @@ CTacFun ([var], arg) - in - let args = List.map map args in - intern_rec env (CAst.make ?loc @@ CTacApp (e, args)) -| CTacApp (f, args) -> - let loc = f.loc in - let (f, ft) = intern_rec env f in - let fold arg (args, t) = - let loc = arg.loc in - let (arg, argt) = intern_rec env arg in - (arg :: args, (loc, argt) :: t) - in - let (args, t) = List.fold_right fold args ([], []) in - let ret = unify_arrow ?loc env ft t in - (GTacApp (f, args), ret) -| CTacLet (is_rec, el, e) -> - let map (pat, e) = - let (pat, ty) = extract_pattern_type pat in - (pat, ty, e) - in - let el = List.map map el in - let fold accu (pat, _, e) = - let ids = ids_of_pattern Id.Set.empty pat in - let common = Id.Set.inter ids accu in - if Id.Set.is_empty common then Id.Set.union ids accu - else - let id = Id.Set.choose common in - user_err ?loc:pat.loc (str "Variable " ++ Id.print id ++ str " is bound several \ - times in this matching") - in - let ids = List.fold_left fold Id.Set.empty el in - if is_rec then intern_let_rec env loc ids el e - else intern_let env loc ids el e -| CTacCnv (e, tc) -> - let (e, t) = intern_rec env e in - let tc = intern_type env tc in - let () = unify ?loc env t tc in - (e, tc) -| CTacSeq (e1, e2) -> - let loc1 = e1.loc in - let (e1, t1) = intern_rec env e1 in - let (e2, t2) = intern_rec env e2 in - let () = check_elt_unit loc1 env t1 in - (GTacLet (false, [Anonymous, e1], e2), t2) -| CTacCse (e, pl) -> - intern_case env loc e pl -| CTacRec fs -> - intern_record env loc fs -| CTacPrj (e, proj) -> - let pinfo = get_projection proj in - let loc = e.loc in - let (e, t) = intern_rec env e in - let subst = Array.init pinfo.pdata_prms (fun _ -> fresh_id env) in - let params = Array.map_to_list (fun i -> GTypVar i) subst in - let exp = GTypRef (Other pinfo.pdata_type, params) in - let () = unify ?loc env t exp in - let substf i = GTypVar subst.(i) in - let ret = subst_type substf pinfo.pdata_ptyp in - (GTacPrj (pinfo.pdata_type, e, pinfo.pdata_indx), ret) -| CTacSet (e, proj, r) -> - let pinfo = get_projection proj in - let () = - if not pinfo.pdata_mutb then - let loc = match proj with - | RelId {CAst.loc} -> loc - | AbsKn _ -> None - in - user_err ?loc (str "Field is not mutable") - in - let subst = Array.init pinfo.pdata_prms (fun _ -> fresh_id env) in - let params = Array.map_to_list (fun i -> GTypVar i) subst in - let exp = GTypRef (Other pinfo.pdata_type, params) in - let e = intern_rec_with_constraint env e exp in - let substf i = GTypVar subst.(i) in - let ret = subst_type substf pinfo.pdata_ptyp in - let r = intern_rec_with_constraint env r ret in - (GTacSet (pinfo.pdata_type, e, pinfo.pdata_indx, r), GTypRef (Tuple 0, [])) -| CTacExt (tag, arg) -> - let open Genintern in - let self ist e = - let env = match Store.get ist.extra ltac2_env with - | None -> empty_env () - | Some env -> env - in - intern_rec env e - in - let obj = interp_ml_object tag in - (* External objects do not have access to the named context because this is - not stable by dynamic semantics. *) - let genv = Global.env_of_context Environ.empty_named_context_val in - let ist = empty_glob_sign genv in - let ist = { ist with extra = Store.set ist.extra ltac2_env env } in - let arg, tpe = - if env.env_str then - let arg () = obj.ml_intern self ist arg in - Flags.with_option Ltac_plugin.Tacintern.strict_check arg () - else - obj.ml_intern self ist arg - in - let e = match arg with - | GlbVal arg -> GTacExt (tag, arg) - | GlbTacexpr e -> e - in - (e, tpe) - -and intern_rec_with_constraint env e exp = - let (er, t) = intern_rec env e in - let () = unify ?loc:e.loc env t exp in - er - -and intern_let env loc ids el e = - let avoid = Id.Set.union ids (Id.Map.domain env.env_var) in - let fold (pat, t, e) (avoid, accu) = - let nas, exp = expand_pattern avoid [pat, t] in - let na = match nas with [x] -> x | _ -> assert false in - let avoid = List.fold_left add_name avoid nas in - (avoid, (na, exp, t, e) :: accu) - in - let (_, el) = List.fold_right fold el (avoid, []) in - let fold (na, exp, tc, e) (body, el, p) = - let (e, t) = match tc with - | None -> intern_rec env e - | Some tc -> - let tc = intern_type env tc in - (intern_rec_with_constraint env e tc, tc) - in - let t = if is_value e then abstract_var env t else monomorphic t in - (exp body, (na, e) :: el, (na, t) :: p) - in - let (e, el, p) = List.fold_right fold el (e, [], []) in - let env = List.fold_left (fun accu (na, t) -> push_name na t accu) env p in - let (e, t) = intern_rec env e in - (GTacLet (false, el, e), t) - -and intern_let_rec env loc ids el e = - let map env (pat, t, e) = - let na = match pat.v with - | CPatVar na -> na - | CPatRef _ | CPatCnv _ -> - user_err ?loc:pat.loc (str "This kind of pattern is forbidden in let-rec bindings") - in - let id = fresh_id env in - let env = push_name na (monomorphic (GTypVar id)) env in - (env, (loc, na, t, e, id)) - in - let (env, el) = List.fold_left_map map env el in - let fold (loc, na, tc, e, id) (el, tl) = - let loc_e = e.loc in - let (e, t) = intern_rec env e in - let () = - if not (is_rec_rhs e) then - user_err ?loc:loc_e (str "This kind of expression is not allowed as \ - right-hand side of a recursive binding") - in - let () = unify ?loc env t (GTypVar id) in - let () = match tc with - | None -> () - | Some tc -> - let tc = intern_type env tc in - unify ?loc env t tc - in - ((na, e) :: el, t :: tl) - in - let (el, tl) = List.fold_right fold el ([], []) in - let (e, t) = intern_rec env e in - (GTacLet (true, el, e), t) - -(** For now, patterns recognized by the pattern-matching compiling are limited - to depth-one where leaves are either variables or catch-all *) -and intern_case env loc e pl = - let (e', t) = intern_rec env e in - let todo ?loc () = user_err ?loc (str "Pattern not handled yet") in - match get_pattern_kind env pl with - | PKind_any -> - let (pat, b) = List.hd pl in - let na = match intern_patexpr env pat with - | GPatVar na -> na - | _ -> assert false - in - let () = check_redundant_clause (List.tl pl) in - let env = push_name na (monomorphic t) env in - let (b, tb) = intern_rec env b in - (GTacLet (false, [na, e'], b), tb) - | PKind_empty -> - let kn = check_elt_empty loc env t in - let r = fresh_id env in - (GTacCse (e', Other kn, [||], [||]), GTypVar r) - | PKind_variant kn -> - let subst, tc = fresh_reftype env kn in - let () = unify ?loc:e.loc env t tc in - let (nconst, nnonconst, arities) = match kn with - | Tuple 0 -> 1, 0, [0] - | Tuple n -> 0, 1, [n] - | Other kn -> - let (_, def) = Tac2env.interp_type kn in - let galg = match def with | GTydAlg c -> c | _ -> assert false in - let arities = List.map (fun (_, args) -> List.length args) galg.galg_constructors in - galg.galg_nconst, galg.galg_nnonconst, arities - in - let const = Array.make nconst None in - let nonconst = Array.make nnonconst None in - let ret = GTypVar (fresh_id env) in - let rec intern_branch = function - | [] -> () - | (pat, br) :: rem -> - let tbr = match pat.v with - | CPatVar (Name _) -> - let loc = pat.loc in - todo ?loc () - | CPatVar Anonymous -> - let () = check_redundant_clause rem in - let (br', brT) = intern_rec env br in - (* Fill all remaining branches *) - let fill (ncst, narg) arity = - if Int.equal arity 0 then - let () = - if Option.is_empty const.(ncst) then const.(ncst) <- Some br' - in - (succ ncst, narg) - else - let () = - if Option.is_empty nonconst.(narg) then - let ids = Array.make arity Anonymous in - nonconst.(narg) <- Some (ids, br') - in - (ncst, succ narg) - in - let _ = List.fold_left fill (0, 0) arities in - brT - | CPatRef (qid, args) -> - let loc = pat.loc in - let knc = get_constructor env qid in - let kn', index, arity = match knc with - | Tuple n -> Tuple n, 0, List.init n (fun i -> GTypVar i) - | Other knc -> - let data = Tac2env.interp_constructor knc in - let index = Option.get data.cdata_indx in - Other data.cdata_type, index, data.cdata_args - in - let () = - if not (eq_or_tuple KerName.equal kn kn') then - invalid_pattern ?loc kn kn' - in - let get_id pat = match pat with - | {v=CPatVar na} -> na - | {loc} -> todo ?loc () - in - let ids = List.map get_id args in - let nids = List.length ids in - let nargs = List.length arity in - let () = match knc with - | Tuple n -> assert (n == nids) - | Other knc -> - if not (Int.equal nids nargs) then error_nargs_mismatch ?loc knc nargs nids - in - let fold env id tpe = - (* Instantiate all arguments *) - let subst n = GTypVar subst.(n) in - let tpe = subst_type subst tpe in - push_name id (monomorphic tpe) env - in - let nenv = List.fold_left2 fold env ids arity in - let (br', brT) = intern_rec nenv br in - let () = - if List.is_empty args then - if Option.is_empty const.(index) then const.(index) <- Some br' - else warn_redundant_clause ?loc () - else - let ids = Array.of_list ids in - if Option.is_empty nonconst.(index) then nonconst.(index) <- Some (ids, br') - else warn_redundant_clause ?loc () - in - brT - | CPatCnv _ -> - user_err ?loc (str "Pattern not handled yet") - in - let () = unify ?loc:br.loc env tbr ret in - intern_branch rem - in - let () = intern_branch pl in - let map n is_const = function - | None -> - let kn = match kn with Other kn -> kn | _ -> assert false in - let cstr = pr_internal_constructor kn n is_const in - user_err ?loc (str "Unhandled match case for constructor " ++ cstr) - | Some x -> x - in - let const = Array.mapi (fun i o -> map i true o) const in - let nonconst = Array.mapi (fun i o -> map i false o) nonconst in - let ce = GTacCse (e', kn, const, nonconst) in - (ce, ret) - | PKind_open kn -> - let subst, tc = fresh_reftype env (Other kn) in - let () = unify ?loc:e.loc env t tc in - let ret = GTypVar (fresh_id env) in - let rec intern_branch map = function - | [] -> - user_err ?loc (str "Missing default case") - | (pat, br) :: rem -> - match intern_patexpr env pat with - | GPatVar na -> - let () = check_redundant_clause rem in - let nenv = push_name na (monomorphic tc) env in - let br' = intern_rec_with_constraint nenv br ret in - let def = (na, br') in - (map, def) - | GPatRef (knc, args) -> - let get = function - | GPatVar na -> na - | GPatRef _ -> - user_err ?loc (str "TODO: Unhandled match case") (* FIXME *) - in - let loc = pat.loc in - let knc = match knc with - | Other knc -> knc - | Tuple n -> invalid_pattern ?loc (Other kn) (Tuple n) - in - let ids = List.map get args in - let data = Tac2env.interp_constructor knc in - let () = - if not (KerName.equal kn data.cdata_type) then - invalid_pattern ?loc (Other kn) (Other data.cdata_type) - in - let nids = List.length ids in - let nargs = List.length data.cdata_args in - let () = - if not (Int.equal nids nargs) then error_nargs_mismatch ?loc knc nargs nids - in - let fold env id tpe = - (* Instantiate all arguments *) - let subst n = GTypVar subst.(n) in - let tpe = subst_type subst tpe in - push_name id (monomorphic tpe) env - in - let nenv = List.fold_left2 fold env ids data.cdata_args in - let br' = intern_rec_with_constraint nenv br ret in - let map = - if KNmap.mem knc map then - let () = warn_redundant_clause ?loc () in - map - else - KNmap.add knc (Anonymous, Array.of_list ids, br') map - in - intern_branch map rem - in - let (map, def) = intern_branch KNmap.empty pl in - (GTacWth { opn_match = e'; opn_branch = map; opn_default = def }, ret) - -and intern_constructor env loc kn args = match kn with -| Other kn -> - let cstr = interp_constructor kn in - let nargs = List.length cstr.cdata_args in - if Int.equal nargs (List.length args) then - let subst = Array.init cstr.cdata_prms (fun _ -> fresh_id env) in - let substf i = GTypVar subst.(i) in - let types = List.map (fun t -> subst_type substf t) cstr.cdata_args in - let targs = List.init cstr.cdata_prms (fun i -> GTypVar subst.(i)) in - let ans = GTypRef (Other cstr.cdata_type, targs) in - let map arg tpe = intern_rec_with_constraint env arg tpe in - let args = List.map2 map args types in - match cstr.cdata_indx with - | Some idx -> - (GTacCst (Other cstr.cdata_type, idx, args), ans) - | None -> - (GTacOpn (kn, args), ans) - else - error_nargs_mismatch ?loc kn nargs (List.length args) -| Tuple n -> - assert (Int.equal n (List.length args)); - let types = List.init n (fun i -> GTypVar (fresh_id env)) in - let map arg tpe = intern_rec_with_constraint env arg tpe in - let args = List.map2 map args types in - let ans = GTypRef (Tuple n, types) in - GTacCst (Tuple n, 0, args), ans - -and intern_record env loc fs = - let map (proj, e) = - let loc = match proj with - | RelId {CAst.loc} -> loc - | AbsKn _ -> None - in - let proj = get_projection proj in - (loc, proj, e) - in - let fs = List.map map fs in - let kn = match fs with - | [] -> user_err ?loc (str "Cannot infer the corresponding record type") - | (_, proj, _) :: _ -> proj.pdata_type - in - let params, typdef = match Tac2env.interp_type kn with - | n, GTydRec def -> n, def - | _ -> assert false - in - let subst = Array.init params (fun _ -> fresh_id env) in - (* Set the answer [args] imperatively *) - let args = Array.make (List.length typdef) None in - let iter (loc, pinfo, e) = - if KerName.equal kn pinfo.pdata_type then - let index = pinfo.pdata_indx in - match args.(index) with - | None -> - let exp = subst_type (fun i -> GTypVar subst.(i)) pinfo.pdata_ptyp in - let e = intern_rec_with_constraint env e exp in - args.(index) <- Some e - | Some _ -> - let (name, _, _) = List.nth typdef pinfo.pdata_indx in - user_err ?loc (str "Field " ++ Id.print name ++ str " is defined \ - several times") - else - user_err ?loc (str "Field " ++ (*KerName.print knp ++*) str " does not \ - pertain to record definition " ++ pr_typref pinfo.pdata_type) - in - let () = List.iter iter fs in - let () = match Array.findi (fun _ o -> Option.is_empty o) args with - | None -> () - | Some i -> - let (field, _, _) = List.nth typdef i in - user_err ?loc (str "Field " ++ Id.print field ++ str " is undefined") - in - let args = Array.map_to_list Option.get args in - let tparam = List.init params (fun i -> GTypVar subst.(i)) in - (GTacCst (Other kn, 0, args), GTypRef (Other kn, tparam)) - -let normalize env (count, vars) (t : UF.elt glb_typexpr) = - let get_var id = - try UF.Map.find id !vars - with Not_found -> - let () = assert env.env_opn in - let n = GTypVar !count in - let () = incr count in - let () = vars := UF.Map.add id n !vars in - n - in - let rec subst id = match UF.find id env.env_cst with - | id, None -> get_var id - | _, Some t -> subst_type subst t - in - subst_type subst t - -let intern ~strict e = - let env = empty_env () in - let env = if strict then env else { env with env_str = false } in - let (e, t) = intern_rec env e in - let count = ref 0 in - let vars = ref UF.Map.empty in - let t = normalize env (count, vars) t in - (e, (!count, t)) - -let intern_typedef self (ids, t) : glb_quant_typedef = - let env = { (empty_env ()) with env_rec = self } in - (* Initialize type parameters *) - let map id = get_alias id env in - let ids = List.map map ids in - let count = ref (List.length ids) in - let vars = ref UF.Map.empty in - let iter n id = vars := UF.Map.add id (GTypVar n) !vars in - let () = List.iteri iter ids in - (* Do not accept unbound type variables *) - let env = { env with env_opn = false } in - let intern t = - let t = intern_type env t in - normalize env (count, vars) t - in - let count = !count in - match t with - | CTydDef None -> (count, GTydDef None) - | CTydDef (Some t) -> (count, GTydDef (Some (intern t))) - | CTydAlg constrs -> - let map (c, t) = (c, List.map intern t) in - let constrs = List.map map constrs in - let getn (const, nonconst) (c, args) = match args with - | [] -> (succ const, nonconst) - | _ :: _ -> (const, succ nonconst) - in - let nconst, nnonconst = List.fold_left getn (0, 0) constrs in - let galg = { - galg_constructors = constrs; - galg_nconst = nconst; - galg_nnonconst = nnonconst; - } in - (count, GTydAlg galg) - | CTydRec fields -> - let map (c, mut, t) = (c, mut, intern t) in - let fields = List.map map fields in - (count, GTydRec fields) - | CTydOpn -> (count, GTydOpn) - -let intern_open_type t = - let env = empty_env () in - let t = intern_type env t in - let count = ref 0 in - let vars = ref UF.Map.empty in - let t = normalize env (count, vars) t in - (!count, t) - -(** Subtyping *) - -let check_subtype t1 t2 = - let env = empty_env () in - let t1 = fresh_type_scheme env t1 in - (* We build a substitution mimicking rigid variable by using dummy tuples *) - let rigid i = GTypRef (Tuple (i + 1), []) in - let (n, t2) = t2 in - let subst = Array.init n rigid in - let substf i = subst.(i) in - let t2 = subst_type substf t2 in - try unify0 env t1 t2; true with CannotUnify _ -> false - -(** Globalization *) - -let get_projection0 var = match var with -| RelId qid -> - let kn = try Tac2env.locate_projection qid with Not_found -> - user_err ?loc:qid.CAst.loc (pr_qualid qid ++ str " is not a projection") - in - kn -| AbsKn kn -> kn - -let rec globalize ids ({loc;v=er} as e) = match er with -| CTacAtm _ -> e -| CTacRef ref -> - let mem id = Id.Set.mem id ids in - begin match get_variable0 mem ref with - | ArgVar _ -> e - | ArgArg kn -> CAst.make ?loc @@ CTacRef (AbsKn kn) - end -| CTacCst qid -> - let knc = get_constructor () qid in - CAst.make ?loc @@ CTacCst (AbsKn knc) -| CTacFun (bnd, e) -> - let fold (pats, accu) pat = - let accu = ids_of_pattern accu pat in - let pat = globalize_pattern ids pat in - (pat :: pats, accu) - in - let bnd, ids = List.fold_left fold ([], ids) bnd in - let bnd = List.rev bnd in - let e = globalize ids e in - CAst.make ?loc @@ CTacFun (bnd, e) -| CTacApp (e, el) -> - let e = globalize ids e in - let el = List.map (fun e -> globalize ids e) el in - CAst.make ?loc @@ CTacApp (e, el) -| CTacLet (isrec, bnd, e) -> - let fold accu (pat, _) = ids_of_pattern accu pat in - let ext = List.fold_left fold Id.Set.empty bnd in - let eids = Id.Set.union ext ids in - let e = globalize eids e in - let map (qid, e) = - let ids = if isrec then eids else ids in - let qid = globalize_pattern ids qid in - (qid, globalize ids e) - in - let bnd = List.map map bnd in - CAst.make ?loc @@ CTacLet (isrec, bnd, e) -| CTacCnv (e, t) -> - let e = globalize ids e in - CAst.make ?loc @@ CTacCnv (e, t) -| CTacSeq (e1, e2) -> - let e1 = globalize ids e1 in - let e2 = globalize ids e2 in - CAst.make ?loc @@ CTacSeq (e1, e2) -| CTacCse (e, bl) -> - let e = globalize ids e in - let bl = List.map (fun b -> globalize_case ids b) bl in - CAst.make ?loc @@ CTacCse (e, bl) -| CTacRec r -> - let map (p, e) = - let p = get_projection0 p in - let e = globalize ids e in - (AbsKn p, e) - in - CAst.make ?loc @@ CTacRec (List.map map r) -| CTacPrj (e, p) -> - let e = globalize ids e in - let p = get_projection0 p in - CAst.make ?loc @@ CTacPrj (e, AbsKn p) -| CTacSet (e, p, e') -> - let e = globalize ids e in - let p = get_projection0 p in - let e' = globalize ids e' in - CAst.make ?loc @@ CTacSet (e, AbsKn p, e') -| CTacExt (tag, arg) -> - let arg = str (Tac2dyn.Arg.repr tag) in - CErrors.user_err ?loc (str "Cannot globalize generic arguments of type" ++ spc () ++ arg) - -and globalize_case ids (p, e) = - (globalize_pattern ids p, globalize ids e) - -and globalize_pattern ids ({loc;v=pr} as p) = match pr with -| CPatVar _ -> p -| CPatRef (cst, pl) -> - let knc = get_constructor () cst in - let cst = AbsKn knc in - let pl = List.map (fun p -> globalize_pattern ids p) pl in - CAst.make ?loc @@ CPatRef (cst, pl) -| CPatCnv (pat, ty) -> - let pat = globalize_pattern ids pat in - CAst.make ?loc @@ CPatCnv (pat, ty) - -(** Kernel substitution *) - -open Mod_subst - -let subst_or_tuple f subst o = match o with -| Tuple _ -> o -| Other v -> - let v' = f subst v in - if v' == v then o else Other v' - -let rec subst_type subst t = match t with -| GTypVar _ -> t -| GTypArrow (t1, t2) -> - let t1' = subst_type subst t1 in - let t2' = subst_type subst t2 in - if t1' == t1 && t2' == t2 then t - else GTypArrow (t1', t2') -| GTypRef (kn, tl) -> - let kn' = subst_or_tuple subst_kn subst kn in - let tl' = List.Smart.map (fun t -> subst_type subst t) tl in - if kn' == kn && tl' == tl then t else GTypRef (kn', tl') - -let rec subst_expr subst e = match e with -| GTacAtm _ | GTacVar _ | GTacPrm _ -> e -| GTacRef kn -> GTacRef (subst_kn subst kn) -| GTacFun (ids, e) -> GTacFun (ids, subst_expr subst e) -| GTacApp (f, args) -> - GTacApp (subst_expr subst f, List.map (fun e -> subst_expr subst e) args) -| GTacLet (r, bs, e) -> - let bs = List.map (fun (na, e) -> (na, subst_expr subst e)) bs in - GTacLet (r, bs, subst_expr subst e) -| GTacCst (t, n, el) as e0 -> - let t' = subst_or_tuple subst_kn subst t in - let el' = List.Smart.map (fun e -> subst_expr subst e) el in - if t' == t && el' == el then e0 else GTacCst (t', n, el') -| GTacCse (e, ci, cse0, cse1) -> - let cse0' = Array.map (fun e -> subst_expr subst e) cse0 in - let cse1' = Array.map (fun (ids, e) -> (ids, subst_expr subst e)) cse1 in - let ci' = subst_or_tuple subst_kn subst ci in - GTacCse (subst_expr subst e, ci', cse0', cse1') -| GTacWth { opn_match = e; opn_branch = br; opn_default = (na, def) } as e0 -> - let e' = subst_expr subst e in - let def' = subst_expr subst def in - let fold kn (self, vars, p) accu = - let kn' = subst_kn subst kn in - let p' = subst_expr subst p in - if kn' == kn && p' == p then accu - else KNmap.add kn' (self, vars, p') (KNmap.remove kn accu) - in - let br' = KNmap.fold fold br br in - if e' == e && br' == br && def' == def then e0 - else GTacWth { opn_match = e'; opn_default = (na, def'); opn_branch = br' } -| GTacPrj (kn, e, p) as e0 -> - let kn' = subst_kn subst kn in - let e' = subst_expr subst e in - if kn' == kn && e' == e then e0 else GTacPrj (kn', e', p) -| GTacSet (kn, e, p, r) as e0 -> - let kn' = subst_kn subst kn in - let e' = subst_expr subst e in - let r' = subst_expr subst r in - if kn' == kn && e' == e && r' == r then e0 else GTacSet (kn', e', p, r') -| GTacExt (tag, arg) -> - let tpe = interp_ml_object tag in - let arg' = tpe.ml_subst subst arg in - if arg' == arg then e else GTacExt (tag, arg') -| GTacOpn (kn, el) as e0 -> - let kn' = subst_kn subst kn in - let el' = List.Smart.map (fun e -> subst_expr subst e) el in - if kn' == kn && el' == el then e0 else GTacOpn (kn', el') - -let subst_typedef subst e = match e with -| GTydDef t -> - let t' = Option.Smart.map (fun t -> subst_type subst t) t in - if t' == t then e else GTydDef t' -| GTydAlg galg -> - let map (c, tl as p) = - let tl' = List.Smart.map (fun t -> subst_type subst t) tl in - if tl' == tl then p else (c, tl') - in - let constrs' = List.Smart.map map galg.galg_constructors in - if constrs' == galg.galg_constructors then e - else GTydAlg { galg with galg_constructors = constrs' } -| GTydRec fields -> - let map (c, mut, t as p) = - let t' = subst_type subst t in - if t' == t then p else (c, mut, t') - in - let fields' = List.Smart.map map fields in - if fields' == fields then e else GTydRec fields' -| GTydOpn -> GTydOpn - -let subst_quant_typedef subst (prm, def as qdef) = - let def' = subst_typedef subst def in - if def' == def then qdef else (prm, def') - -let subst_type_scheme subst (prm, t as sch) = - let t' = subst_type subst t in - if t' == t then sch else (prm, t') - -let subst_or_relid subst ref = match ref with -| RelId _ -> ref -| AbsKn kn -> - let kn' = subst_or_tuple subst_kn subst kn in - if kn' == kn then ref else AbsKn kn' - -let rec subst_rawtype subst ({loc;v=tr} as t) = match tr with -| CTypVar _ -> t -| CTypArrow (t1, t2) -> - let t1' = subst_rawtype subst t1 in - let t2' = subst_rawtype subst t2 in - if t1' == t1 && t2' == t2 then t else CAst.make ?loc @@ CTypArrow (t1', t2') -| CTypRef (ref, tl) -> - let ref' = subst_or_relid subst ref in - let tl' = List.Smart.map (fun t -> subst_rawtype subst t) tl in - if ref' == ref && tl' == tl then t else CAst.make ?loc @@ CTypRef (ref', tl') - -let subst_tacref subst ref = match ref with -| RelId _ -> ref -| AbsKn (TacConstant kn) -> - let kn' = subst_kn subst kn in - if kn' == kn then ref else AbsKn (TacConstant kn') -| AbsKn (TacAlias kn) -> - let kn' = subst_kn subst kn in - if kn' == kn then ref else AbsKn (TacAlias kn') - -let subst_projection subst prj = match prj with -| RelId _ -> prj -| AbsKn kn -> - let kn' = subst_kn subst kn in - if kn' == kn then prj else AbsKn kn' - -let rec subst_rawpattern subst ({loc;v=pr} as p) = match pr with -| CPatVar _ -> p -| CPatRef (c, pl) -> - let pl' = List.Smart.map (fun p -> subst_rawpattern subst p) pl in - let c' = subst_or_relid subst c in - if pl' == pl && c' == c then p else CAst.make ?loc @@ CPatRef (c', pl') -| CPatCnv (pat, ty) -> - let pat' = subst_rawpattern subst pat in - let ty' = subst_rawtype subst ty in - if pat' == pat && ty' == ty then p else CAst.make ?loc @@ CPatCnv (pat', ty') - -(** Used for notations *) -let rec subst_rawexpr subst ({loc;v=tr} as t) = match tr with -| CTacAtm _ -> t -| CTacRef ref -> - let ref' = subst_tacref subst ref in - if ref' == ref then t else CAst.make ?loc @@ CTacRef ref' -| CTacCst ref -> - let ref' = subst_or_relid subst ref in - if ref' == ref then t else CAst.make ?loc @@ CTacCst ref' -| CTacFun (bnd, e) -> - let map pat = subst_rawpattern subst pat in - let bnd' = List.Smart.map map bnd in - let e' = subst_rawexpr subst e in - if bnd' == bnd && e' == e then t else CAst.make ?loc @@ CTacFun (bnd', e') -| CTacApp (e, el) -> - let e' = subst_rawexpr subst e in - let el' = List.Smart.map (fun e -> subst_rawexpr subst e) el in - if e' == e && el' == el then t else CAst.make ?loc @@ CTacApp (e', el') -| CTacLet (isrec, bnd, e) -> - let map (na, e as p) = - let na' = subst_rawpattern subst na in - let e' = subst_rawexpr subst e in - if na' == na && e' == e then p else (na', e') - in - let bnd' = List.Smart.map map bnd in - let e' = subst_rawexpr subst e in - if bnd' == bnd && e' == e then t else CAst.make ?loc @@ CTacLet (isrec, bnd', e') -| CTacCnv (e, c) -> - let e' = subst_rawexpr subst e in - let c' = subst_rawtype subst c in - if c' == c && e' == e then t else CAst.make ?loc @@ CTacCnv (e', c') -| CTacSeq (e1, e2) -> - let e1' = subst_rawexpr subst e1 in - let e2' = subst_rawexpr subst e2 in - if e1' == e1 && e2' == e2 then t else CAst.make ?loc @@ CTacSeq (e1', e2') -| CTacCse (e, bl) -> - let map (p, e as x) = - let p' = subst_rawpattern subst p in - let e' = subst_rawexpr subst e in - if p' == p && e' == e then x else (p', e') - in - let e' = subst_rawexpr subst e in - let bl' = List.Smart.map map bl in - if e' == e && bl' == bl then t else CAst.make ?loc @@ CTacCse (e', bl') -| CTacRec el -> - let map (prj, e as p) = - let prj' = subst_projection subst prj in - let e' = subst_rawexpr subst e in - if prj' == prj && e' == e then p else (prj', e') - in - let el' = List.Smart.map map el in - if el' == el then t else CAst.make ?loc @@ CTacRec el' -| CTacPrj (e, prj) -> - let prj' = subst_projection subst prj in - let e' = subst_rawexpr subst e in - if prj' == prj && e' == e then t else CAst.make ?loc @@ CTacPrj (e', prj') -| CTacSet (e, prj, r) -> - let prj' = subst_projection subst prj in - let e' = subst_rawexpr subst e in - let r' = subst_rawexpr subst r in - if prj' == prj && e' == e && r' == r then t else CAst.make ?loc @@ CTacSet (e', prj', r') -| CTacExt _ -> assert false (** Should not be generated by globalization *) - -(** Registering *) - -let () = - let open Genintern in - let intern ist tac = - let env = match Genintern.Store.get ist.extra ltac2_env with - | None -> - (* Only happens when Ltac2 is called from a constr or ltac1 quotation *) - let env = empty_env () in - if !Ltac_plugin.Tacintern.strict_check then env - else { env with env_str = false } - | Some env -> env - in - let loc = tac.loc in - let (tac, t) = intern_rec env tac in - let () = check_elt_unit loc env t in - (ist, tac) - in - Genintern.register_intern0 wit_ltac2 intern -let () = Genintern.register_subst0 wit_ltac2 subst_expr - -let () = - let open Genintern in - let intern ist (loc, id) = - let env = match Genintern.Store.get ist.extra ltac2_env with - | None -> - (* Only happens when Ltac2 is called from a constr or ltac1 quotation *) - let env = empty_env () in - if !Ltac_plugin.Tacintern.strict_check then env - else { env with env_str = false } - | Some env -> env - in - let t = - try Id.Map.find id env.env_var - with Not_found -> - CErrors.user_err ?loc (str "Unbound value " ++ Id.print id) - in - let t = fresh_mix_type_scheme env t in - let () = unify ?loc env t (GTypRef (Other t_constr, [])) in - (ist, id) - in - Genintern.register_intern0 wit_ltac2_quotation intern - -let () = Genintern.register_subst0 wit_ltac2_quotation (fun _ id -> id) diff --git a/src/tac2intern.mli b/src/tac2intern.mli deleted file mode 100644 index d646b5cda5..0000000000 --- a/src/tac2intern.mli +++ /dev/null @@ -1,46 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* raw_tacexpr -> glb_tacexpr * type_scheme -val intern_typedef : (KerName.t * int) Id.Map.t -> raw_quant_typedef -> glb_quant_typedef -val intern_open_type : raw_typexpr -> type_scheme - -(** Check that a term is a value. Only values are safe to marshall between - processes. *) -val is_value : glb_tacexpr -> bool -val check_unit : ?loc:Loc.t -> type_scheme -> unit - -val check_subtype : type_scheme -> type_scheme -> bool -(** [check_subtype t1 t2] returns [true] iff all values of intances of type [t1] - also have type [t2]. *) - -val subst_type : substitution -> 'a glb_typexpr -> 'a glb_typexpr -val subst_expr : substitution -> glb_tacexpr -> glb_tacexpr -val subst_quant_typedef : substitution -> glb_quant_typedef -> glb_quant_typedef -val subst_type_scheme : substitution -> type_scheme -> type_scheme - -val subst_rawexpr : substitution -> raw_tacexpr -> raw_tacexpr - -(** {5 Notations} *) - -val globalize : Id.Set.t -> raw_tacexpr -> raw_tacexpr -(** Replaces all qualified identifiers by their corresponding kernel name. The - set represents bound variables in the context. *) - -(** Errors *) - -val error_nargs_mismatch : ?loc:Loc.t -> ltac_constructor -> int -> int -> 'a -val error_nparams_mismatch : ?loc:Loc.t -> int -> int -> 'a - -(** Misc *) - -val drop_ltac2_env : Genintern.Store.t -> Genintern.Store.t diff --git a/src/tac2interp.ml b/src/tac2interp.ml deleted file mode 100644 index b0f8083aeb..0000000000 --- a/src/tac2interp.ml +++ /dev/null @@ -1,227 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* >= fun sigma -> - match Evd.Store.get (Evd.get_extra_data sigma) backtrace with - | None -> Proofview.tclUNIT [] - | Some bt -> Proofview.tclUNIT bt - -let set_backtrace bt = - Proofview.tclEVARMAP >>= fun sigma -> - let store = Evd.get_extra_data sigma in - let store = Evd.Store.set store backtrace bt in - let sigma = Evd.set_extra_data store sigma in - Proofview.Unsafe.tclEVARS sigma - -let with_frame frame tac = - if !print_ltac2_backtrace then - get_backtrace >>= fun bt -> - set_backtrace (frame :: bt) >>= fun () -> - tac >>= fun ans -> - set_backtrace bt >>= fun () -> - Proofview.tclUNIT ans - else tac - -type environment = Tac2env.environment = { - env_ist : valexpr Id.Map.t; -} - -let empty_environment = { - env_ist = Id.Map.empty; -} - -type closure = { - mutable clos_env : valexpr Id.Map.t; - (** Mutable so that we can implement recursive functions imperatively *) - clos_var : Name.t list; - (** Bound variables *) - clos_exp : glb_tacexpr; - (** Body *) - clos_ref : ltac_constant option; - (** Global constant from which the closure originates *) -} - -let push_name ist id v = match id with -| Anonymous -> ist -| Name id -> { env_ist = Id.Map.add id v ist.env_ist } - -let get_var ist id = - try Id.Map.find id ist.env_ist with Not_found -> - anomaly (str "Unbound variable " ++ Id.print id) - -let get_ref ist kn = - try - let data = Tac2env.interp_global kn in - data.Tac2env.gdata_expr - with Not_found -> - anomaly (str "Unbound reference" ++ KerName.print kn) - -let return = Proofview.tclUNIT - -let rec interp (ist : environment) = function -| GTacAtm (AtmInt n) -> return (Tac2ffi.of_int n) -| GTacAtm (AtmStr s) -> return (Tac2ffi.of_string (Bytes.of_string s)) -| GTacVar id -> return (get_var ist id) -| GTacRef kn -> - let data = get_ref ist kn in - return (eval_pure (Some kn) data) -| GTacFun (ids, e) -> - let cls = { clos_ref = None; clos_env = ist.env_ist; clos_var = ids; clos_exp = e } in - let f = interp_app cls in - return (Tac2ffi.of_closure f) -| GTacApp (f, args) -> - interp ist f >>= fun f -> - Proofview.Monad.List.map (fun e -> interp ist e) args >>= fun args -> - Tac2ffi.apply (Tac2ffi.to_closure f) args -| GTacLet (false, el, e) -> - let fold accu (na, e) = - interp ist e >>= fun e -> - return (push_name accu na e) - in - Proofview.Monad.List.fold_left fold ist el >>= fun ist -> - interp ist e -| GTacLet (true, el, e) -> - let map (na, e) = match e with - | GTacFun (ids, e) -> - let cls = { clos_ref = None; clos_env = ist.env_ist; clos_var = ids; clos_exp = e } in - let f = Tac2ffi.of_closure (interp_app cls) in - na, cls, f - | _ -> anomaly (str "Ill-formed recursive function") - in - let fixs = List.map map el in - let fold accu (na, _, cls) = match na with - | Anonymous -> accu - | Name id -> { env_ist = Id.Map.add id cls accu.env_ist } - in - let ist = List.fold_left fold ist fixs in - (* Hack to make a cycle imperatively in the environment *) - let iter (_, e, _) = e.clos_env <- ist.env_ist in - let () = List.iter iter fixs in - interp ist e -| GTacCst (_, n, []) -> return (Valexpr.make_int n) -| GTacCst (_, n, el) -> - Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> - return (Valexpr.make_block n (Array.of_list el)) -| GTacCse (e, _, cse0, cse1) -> - interp ist e >>= fun e -> interp_case ist e cse0 cse1 -| GTacWth { opn_match = e; opn_branch = cse; opn_default = def } -> - interp ist e >>= fun e -> interp_with ist e cse def -| GTacPrj (_, e, p) -> - interp ist e >>= fun e -> interp_proj ist e p -| GTacSet (_, e, p, r) -> - interp ist e >>= fun e -> - interp ist r >>= fun r -> - interp_set ist e p r -| GTacOpn (kn, el) -> - Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> - return (Tac2ffi.of_open (kn, Array.of_list el)) -| GTacPrm (ml, el) -> - Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> - with_frame (FrPrim ml) (Tac2ffi.apply (Tac2env.interp_primitive ml) el) -| GTacExt (tag, e) -> - let tpe = Tac2env.interp_ml_object tag in - with_frame (FrExtn (tag, e)) (tpe.Tac2env.ml_interp ist e) - -and interp_app f = - let ans = fun args -> - let { clos_env = ist; clos_var = ids; clos_exp = e; clos_ref = kn } = f in - let frame = match kn with - | None -> FrAnon e - | Some kn -> FrLtac kn - in - let ist = { env_ist = ist } in - let ist = List.fold_left2 push_name ist ids args in - with_frame frame (interp ist e) - in - Tac2ffi.abstract (List.length f.clos_var) ans - -and interp_case ist e cse0 cse1 = - if Valexpr.is_int e then - interp ist cse0.(Tac2ffi.to_int e) - else - let (n, args) = Tac2ffi.to_block e in - let (ids, e) = cse1.(n) in - let ist = CArray.fold_left2 push_name ist ids args in - interp ist e - -and interp_with ist e cse def = - let (kn, args) = Tac2ffi.to_open e in - let br = try Some (KNmap.find kn cse) with Not_found -> None in - begin match br with - | None -> - let (self, def) = def in - let ist = push_name ist self e in - interp ist def - | Some (self, ids, p) -> - let ist = push_name ist self e in - let ist = CArray.fold_left2 push_name ist ids args in - interp ist p - end - -and interp_proj ist e p = - return (Valexpr.field e p) - -and interp_set ist e p r = - let () = Valexpr.set_field e p r in - return (Valexpr.make_int 0) - -and eval_pure kn = function -| GTacAtm (AtmInt n) -> Valexpr.make_int n -| GTacRef kn -> - let { Tac2env.gdata_expr = e } = - try Tac2env.interp_global kn - with Not_found -> assert false - in - eval_pure (Some kn) e -| GTacFun (na, e) -> - let cls = { clos_ref = kn; clos_env = Id.Map.empty; clos_var = na; clos_exp = e } in - let f = interp_app cls in - Tac2ffi.of_closure f -| GTacCst (_, n, []) -> Valexpr.make_int n -| GTacCst (_, n, el) -> Valexpr.make_block n (Array.map_of_list eval_unnamed el) -| GTacOpn (kn, el) -> Tac2ffi.of_open (kn, Array.map_of_list eval_unnamed el) -| GTacAtm (AtmStr _) | GTacLet _ | GTacVar _ | GTacSet _ -| GTacApp _ | GTacCse _ | GTacPrj _ | GTacPrm _ | GTacExt _ | GTacWth _ -> - anomaly (Pp.str "Term is not a syntactical value") - -and eval_unnamed e = eval_pure None e - - -(** Cross-boundary hacks. *) - -open Geninterp - -let val_env : environment Val.typ = Val.create "ltac2:env" -let env_ref = Id.of_string_soft "@@ltac2_env@@" - -let extract_env (Val.Dyn (tag, v)) : environment = -match Val.eq tag val_env with -| None -> assert false -| Some Refl -> v - -let get_env ist = - try extract_env (Id.Map.find env_ref ist) - with Not_found -> empty_environment - -let set_env env ist = - Id.Map.add env_ref (Val.Dyn (val_env, env)) ist diff --git a/src/tac2interp.mli b/src/tac2interp.mli deleted file mode 100644 index 21fdcd03af..0000000000 --- a/src/tac2interp.mli +++ /dev/null @@ -1,37 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* glb_tacexpr -> valexpr Proofview.tactic - -(* val interp_app : closure -> ml_tactic *) - -(** {5 Cross-boundary encodings} *) - -val get_env : Ltac_pretype.unbound_ltac_var_map -> environment -val set_env : environment -> Ltac_pretype.unbound_ltac_var_map -> Ltac_pretype.unbound_ltac_var_map - -(** {5 Exceptions} *) - -exception LtacError of KerName.t * valexpr array -(** Ltac2-defined exceptions seen from OCaml side *) - -(** {5 Backtrace} *) - -val get_backtrace : backtrace Proofview.tactic - -val with_frame : frame -> 'a Proofview.tactic -> 'a Proofview.tactic - -val print_ltac2_backtrace : bool ref diff --git a/src/tac2match.ml b/src/tac2match.ml deleted file mode 100644 index c9e549d47e..0000000000 --- a/src/tac2match.ml +++ /dev/null @@ -1,232 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* None - | None, Some c | Some c, None -> Some c - | Some c1, Some c2 -> - if equal_instances env sigma c1 c2 then Some c1 - else raise Not_coherent_metas - in - Id.Map.merge merge s1 s2 - -let matching_error = - CErrors.UserError (Some "tactic matching" , Pp.str "No matching clauses for match.") - -let imatching_error = (matching_error, Exninfo.null) - -(** A functor is introduced to share the environment and the - evar_map. They do not change and it would be a pity to introduce - closures everywhere just for the occasional calls to - {!equal_instances}. *) -module type StaticEnvironment = sig - val env : Environ.env - val sigma : Evd.evar_map -end -module PatternMatching (E:StaticEnvironment) = struct - - - (** {6 The pattern-matching monad } *) - - - (** To focus on the algorithmic portion of pattern-matching, the - bookkeeping is relegated to a monad: the composition of the - bactracking monad of {!IStream.t} with a "writer" effect. *) - (* spiwack: as we don't benefit from the various stream optimisations - of Haskell, it may be costly to give the monad in direct style such as - here. We may want to use some continuation passing style. *) - type 'a tac = 'a Proofview.tactic - type 'a m = { stream : 'r. ('a -> result -> 'r tac) -> result -> 'r tac } - - (** The empty substitution. *) - let empty_subst = Id.Map.empty - - (** Composes two substitutions using {!verify_metas_coherence}. It - must be a monoid with neutral element {!empty_subst}. Raises - [Not_coherent_metas] when composition cannot be achieved. *) - let subst_prod s1 s2 = - if is_empty_subst s1 then s2 - else if is_empty_subst s2 then s1 - else verify_metas_coherence E.env E.sigma s1 s2 - - (** Merge two writers (and ignore the first value component). *) - let merge m1 m2 = - try Some { - subst = subst_prod m1.subst m2.subst; - } - with Not_coherent_metas -> None - - (** Monadic [return]: returns a single success with empty substitutions. *) - let return (type a) (lhs:a) : a m = - { stream = fun k ctx -> k lhs ctx } - - (** Monadic bind: each success of [x] is replaced by the successes - of [f x]. The substitutions of [x] and [f x] are composed, - dropping the apparent successes when the substitutions are not - coherent. *) - let (>>=) (type a) (type b) (m:a m) (f:a -> b m) : b m = - { stream = fun k ctx -> m.stream (fun x ctx -> (f x).stream k ctx) ctx } - - (** A variant of [(>>=)] when the first argument returns [unit]. *) - let (<*>) (type a) (m:unit m) (y:a m) : a m = - { stream = fun k ctx -> m.stream (fun () ctx -> y.stream k ctx) ctx } - - (** Failure of the pattern-matching monad: no success. *) - let fail (type a) : a m = { stream = fun _ _ -> Proofview.tclZERO matching_error } - - let run (m : 'a m) = - let ctx = { - subst = empty_subst ; - } in - let eval x ctx = Proofview.tclUNIT (x, ctx) in - m.stream eval ctx - - (** Chooses in a list, in the same order as the list *) - let rec pick (l:'a list) (e, info) : 'a m = match l with - | [] -> { stream = fun _ _ -> Proofview.tclZERO ~info e } - | x :: l -> - { stream = fun k ctx -> Proofview.tclOR (k x ctx) (fun e -> (pick l e).stream k ctx) } - - let pick l = pick l imatching_error - - let put_subst subst : unit m = - let s = { subst } in - { stream = fun k ctx -> match merge s ctx with None -> Proofview.tclZERO matching_error | Some s -> k () s } - - (** {6 Pattern-matching} *) - - let pattern_match_term pat term = - match pat with - | MatchPattern p -> - begin - try - put_subst (Constr_matching.matches E.env E.sigma p term) <*> - return None - with Constr_matching.PatternMatchingFailure -> fail - end - | MatchContext p -> - - let rec map s (e, info) = - { stream = fun k ctx -> match IStream.peek s with - | IStream.Nil -> Proofview.tclZERO ~info e - | IStream.Cons ({ Constr_matching.m_sub = (_, subst); m_ctx }, s) -> - let nctx = { subst } in - match merge ctx nctx with - | None -> (map s (e, info)).stream k ctx - | Some nctx -> Proofview.tclOR (k (Some (Lazy.force m_ctx)) nctx) (fun e -> (map s e).stream k ctx) - } - in - map (Constr_matching.match_subterm E.env E.sigma (Id.Set.empty,p) term) imatching_error - - let hyp_match_type pat hyps = - pick hyps >>= fun decl -> - let id = NamedDecl.get_id decl in - pattern_match_term pat (NamedDecl.get_type decl) >>= fun ctx -> - return (id, ctx) - - let _hyp_match_body_and_type bodypat typepat hyps = - pick hyps >>= function - | LocalDef (id,body,hyp) -> - pattern_match_term bodypat body >>= fun ctx_body -> - pattern_match_term typepat hyp >>= fun ctx_typ -> - return (id, ctx_body, ctx_typ) - | LocalAssum (id,hyp) -> fail - - let hyp_match pat hyps = - match pat with - | typepat -> - hyp_match_type typepat hyps -(* | Def ((_,hypname),bodypat,typepat) -> *) -(* hyp_match_body_and_type hypname bodypat typepat hyps *) - - (** [hyp_pattern_list_match pats hyps lhs], matches the list of - patterns [pats] against the hypotheses in [hyps], and eventually - returns [lhs]. *) - let rec hyp_pattern_list_match pats hyps accu = - match pats with - | pat::pats -> - hyp_match pat hyps >>= fun (matched_hyp, hyp_ctx) -> - let select_matched_hyp decl = Id.equal (NamedDecl.get_id decl) matched_hyp in - let hyps = CList.remove_first select_matched_hyp hyps in - hyp_pattern_list_match pats hyps ((matched_hyp, hyp_ctx) :: accu) - | [] -> return accu - - let rule_match_goal hyps concl = function - | (hyppats,conclpat) -> - (* the rules are applied from the topmost one (in the concrete - syntax) to the bottommost. *) - let hyppats = List.rev hyppats in - pattern_match_term conclpat concl >>= fun ctx_concl -> - hyp_pattern_list_match hyppats hyps [] >>= fun hyps -> - return (hyps, ctx_concl) - -end - -let match_goal env sigma concl ~rev rule = - let open Proofview.Notations in - let hyps = EConstr.named_context env in - let hyps = if rev then List.rev hyps else hyps in - let module E = struct - let env = env - let sigma = sigma - end in - let module M = PatternMatching(E) in - M.run (M.rule_match_goal hyps concl rule) >>= fun ((hyps, ctx_concl), subst) -> - Proofview.tclUNIT (hyps, ctx_concl, subst.subst) diff --git a/src/tac2match.mli b/src/tac2match.mli deleted file mode 100644 index c82c40d238..0000000000 --- a/src/tac2match.mli +++ /dev/null @@ -1,33 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* - Evd.evar_map -> - constr -> - rev:bool -> - match_rule -> - ((Id.t * context option) list * (* List of hypotheses matching: name + context *) - context option * (* Context for conclusion *) - Ltac_pretype.patvar_map (* Pattern variable substitution *)) Proofview.tactic diff --git a/src/tac2print.ml b/src/tac2print.ml deleted file mode 100644 index f4cb290265..0000000000 --- a/src/tac2print.ml +++ /dev/null @@ -1,488 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* str "'" ++ str (pr n) - | GTypRef (Other kn, []) -> pr_typref kn - | GTypRef (Other kn, [t]) -> - let paren = match lvl with - | T5_r | T5_l | T2 | T1 -> fun x -> x - | T0 -> paren - in - paren (pr_glbtype T1 t ++ spc () ++ pr_typref kn) - | GTypRef (Other kn, tl) -> - let paren = match lvl with - | T5_r | T5_l | T2 | T1 -> fun x -> x - | T0 -> paren - in - paren (str "(" ++ prlist_with_sep (fun () -> str ", ") (pr_glbtype lvl) tl ++ str ")" ++ spc () ++ pr_typref kn) - | GTypArrow (t1, t2) -> - let paren = match lvl with - | T5_r -> fun x -> x - | T5_l | T2 | T1 | T0 -> paren - in - paren (pr_glbtype T5_l t1 ++ spc () ++ str "->" ++ spc () ++ pr_glbtype T5_r t2) - | GTypRef (Tuple 0, []) -> - Libnames.pr_qualid (Tac2env.shortest_qualid_of_type t_unit) - | GTypRef (Tuple _, tl) -> - let paren = match lvl with - | T5_r | T5_l -> fun x -> x - | T2 | T1 | T0 -> paren - in - paren (prlist_with_sep (fun () -> str " * ") (pr_glbtype T2) tl) - in - hov 0 (pr_glbtype lvl c) - -let pr_glbtype pr c = pr_glbtype_gen pr T5_r c - -let int_name () = - let vars = ref Int.Map.empty in - fun n -> - if Int.Map.mem n !vars then Int.Map.find n !vars - else - let num = Int.Map.cardinal !vars in - let base = num mod 26 in - let rem = num / 26 in - let name = String.make 1 (Char.chr (97 + base)) in - let suff = if Int.equal rem 0 then "" else string_of_int rem in - let name = name ^ suff in - let () = vars := Int.Map.add n name !vars in - name - -(** Term printing *) - -let pr_constructor kn = - Libnames.pr_qualid (Tac2env.shortest_qualid_of_constructor kn) - -let pr_projection kn = - Libnames.pr_qualid (Tac2env.shortest_qualid_of_projection kn) - -type exp_level = Tac2expr.exp_level = -| E5 -| E4 -| E3 -| E2 -| E1 -| E0 - -let pr_atom = function -| AtmInt n -> Pp.int n -| AtmStr s -> qstring s - -let pr_name = function -| Name id -> Id.print id -| Anonymous -> str "_" - -let find_constructor n empty def = - let rec find n = function - | [] -> assert false - | (id, []) as ans :: rem -> - if empty then - if Int.equal n 0 then ans - else find (pred n) rem - else find n rem - | (id, _ :: _) as ans :: rem -> - if not empty then - if Int.equal n 0 then ans - else find (pred n) rem - else find n rem - in - find n def - -let pr_internal_constructor tpe n is_const = - let data = match Tac2env.interp_type tpe with - | (_, GTydAlg data) -> data - | _ -> assert false - in - let (id, _) = find_constructor n is_const data.galg_constructors in - let kn = change_kn_label tpe id in - pr_constructor kn - -let order_branches cbr nbr def = - let rec order cidx nidx def = match def with - | [] -> [] - | (id, []) :: rem -> - let ans = order (succ cidx) nidx rem in - (id, [], cbr.(cidx)) :: ans - | (id, _ :: _) :: rem -> - let ans = order cidx (succ nidx) rem in - let (vars, e) = nbr.(nidx) in - (id, Array.to_list vars, e) :: ans - in - order 0 0 def - -let pr_glbexpr_gen lvl c = - let rec pr_glbexpr lvl = function - | GTacAtm atm -> pr_atom atm - | GTacVar id -> Id.print id - | GTacRef gr -> - let qid = shortest_qualid_of_ltac (TacConstant gr) in - Libnames.pr_qualid qid - | GTacFun (nas, c) -> - let nas = pr_sequence pr_name nas in - let paren = match lvl with - | E0 | E1 | E2 | E3 | E4 -> paren - | E5 -> fun x -> x - in - paren (hov 0 (hov 2 (str "fun" ++ spc () ++ nas) ++ spc () ++ str "=>" ++ spc () ++ - pr_glbexpr E5 c)) - | GTacApp (c, cl) -> - let paren = match lvl with - | E0 -> paren - | E1 | E2 | E3 | E4 | E5 -> fun x -> x - in - paren (hov 2 (pr_glbexpr E1 c ++ spc () ++ (pr_sequence (pr_glbexpr E0) cl))) - | GTacLet (mut, bnd, e) -> - let paren = match lvl with - | E0 | E1 | E2 | E3 | E4 -> paren - | E5 -> fun x -> x - in - let mut = if mut then str "rec" ++ spc () else mt () in - let pr_bnd (na, e) = - pr_name na ++ spc () ++ str ":=" ++ spc () ++ hov 2 (pr_glbexpr E5 e) ++ spc () - in - let bnd = prlist_with_sep (fun () -> str "with" ++ spc ()) pr_bnd bnd in - paren (hv 0 (hov 2 (str "let" ++ spc () ++ mut ++ bnd ++ str "in") ++ spc () ++ pr_glbexpr E5 e)) - | GTacCst (Tuple 0, _, _) -> str "()" - | GTacCst (Tuple _, _, cl) -> - let paren = match lvl with - | E0 | E1 -> paren - | E2 | E3 | E4 | E5 -> fun x -> x - in - paren (prlist_with_sep (fun () -> str "," ++ spc ()) (pr_glbexpr E1) cl) - | GTacCst (Other tpe, n, cl) -> - pr_applied_constructor lvl tpe n cl - | GTacCse (e, info, cst_br, ncst_br) -> - let e = pr_glbexpr E5 e in - let br = match info with - | Other kn -> - let def = match Tac2env.interp_type kn with - | _, GTydAlg { galg_constructors = def } -> def - | _, GTydDef _ | _, GTydRec _ | _, GTydOpn -> assert false - in - let br = order_branches cst_br ncst_br def in - let pr_branch (cstr, vars, p) = - let cstr = change_kn_label kn cstr in - let cstr = pr_constructor cstr in - let vars = match vars with - | [] -> mt () - | _ -> spc () ++ pr_sequence pr_name vars - in - hov 4 (str "|" ++ spc () ++ hov 0 (cstr ++ vars ++ spc () ++ str "=>") ++ spc () ++ - hov 2 (pr_glbexpr E5 p)) ++ spc () - in - prlist pr_branch br - | Tuple n -> - let (vars, p) = if Int.equal n 0 then ([||], cst_br.(0)) else ncst_br.(0) in - let p = pr_glbexpr E5 p in - let vars = prvect_with_sep (fun () -> str "," ++ spc ()) pr_name vars in - hov 4 (str "|" ++ spc () ++ hov 0 (paren vars ++ spc () ++ str "=>") ++ spc () ++ p) - in - v 0 (hv 0 (str "match" ++ spc () ++ e ++ spc () ++ str "with") ++ spc () ++ br ++ spc () ++ str "end") - | GTacWth wth -> - let e = pr_glbexpr E5 wth.opn_match in - let pr_pattern c self vars p = - let self = match self with - | Anonymous -> mt () - | Name id -> spc () ++ str "as" ++ spc () ++ Id.print id - in - hov 4 (str "|" ++ spc () ++ hov 0 (c ++ vars ++ self ++ spc () ++ str "=>") ++ spc () ++ - hov 2 (pr_glbexpr E5 p)) ++ spc () - in - let pr_branch (cstr, (self, vars, p)) = - let cstr = pr_constructor cstr in - let vars = match Array.to_list vars with - | [] -> mt () - | vars -> spc () ++ pr_sequence pr_name vars - in - pr_pattern cstr self vars p - in - let br = prlist pr_branch (KNmap.bindings wth.opn_branch) in - let (def_as, def_p) = wth.opn_default in - let def = pr_pattern (str "_") def_as (mt ()) def_p in - let br = br ++ def in - v 0 (hv 0 (str "match" ++ spc () ++ e ++ spc () ++ str "with") ++ spc () ++ br ++ str "end") - | GTacPrj (kn, e, n) -> - let def = match Tac2env.interp_type kn with - | _, GTydRec def -> def - | _, GTydDef _ | _, GTydAlg _ | _, GTydOpn -> assert false - in - let (proj, _, _) = List.nth def n in - let proj = change_kn_label kn proj in - let proj = pr_projection proj in - let e = pr_glbexpr E0 e in - hov 0 (e ++ str "." ++ paren proj) - | GTacSet (kn, e, n, r) -> - let def = match Tac2env.interp_type kn with - | _, GTydRec def -> def - | _, GTydDef _ | _, GTydAlg _ | _, GTydOpn -> assert false - in - let (proj, _, _) = List.nth def n in - let proj = change_kn_label kn proj in - let proj = pr_projection proj in - let e = pr_glbexpr E0 e in - let r = pr_glbexpr E1 r in - hov 0 (e ++ str "." ++ paren proj ++ spc () ++ str ":=" ++ spc () ++ r) - | GTacOpn (kn, cl) -> - let paren = match lvl with - | E0 -> paren - | E1 | E2 | E3 | E4 | E5 -> fun x -> x - in - let c = pr_constructor kn in - paren (hov 0 (c ++ spc () ++ (pr_sequence (pr_glbexpr E0) cl))) - | GTacExt (tag, arg) -> - let tpe = interp_ml_object tag in - hov 0 (tpe.ml_print (Global.env ()) arg) (* FIXME *) - | GTacPrm (prm, args) -> - let args = match args with - | [] -> mt () - | _ -> spc () ++ pr_sequence (pr_glbexpr E0) args - in - hov 0 (str "@external" ++ spc () ++ qstring prm.mltac_plugin ++ spc () ++ - qstring prm.mltac_tactic ++ args) - and pr_applied_constructor lvl tpe n cl = - let _, data = Tac2env.interp_type tpe in - if KerName.equal tpe t_list then - let rec factorize accu = function - | GTacCst (_, 0, []) -> accu, None - | GTacCst (_, 0, [e; l]) -> factorize (e :: accu) l - | e -> accu, Some e - in - let l, e = factorize [] (GTacCst (Other tpe, n, cl)) in - match e with - | None -> - let pr e = pr_glbexpr E4 e in - hov 2 (str "[" ++ prlist_with_sep pr_semicolon pr (List.rev l) ++ str "]") - | Some e -> - let paren = match lvl with - | E0 | E1 | E2 -> paren - | E3 | E4 | E5 -> fun x -> x - in - let pr e = pr_glbexpr E1 e in - let pr_cons () = spc () ++ str "::" ++ spc () in - paren (hov 2 (prlist_with_sep pr_cons pr (List.rev (e :: l)))) - else match data with - | GTydAlg def -> - let paren = match lvl with - | E0 -> - if List.is_empty cl then fun x -> x else paren - | E1 | E2 | E3 | E4 | E5 -> fun x -> x - in - let cstr = pr_internal_constructor tpe n (List.is_empty cl) in - let cl = match cl with - | [] -> mt () - | _ -> spc () ++ pr_sequence (pr_glbexpr E0) cl - in - paren (hov 2 (cstr ++ cl)) - | GTydRec def -> - let args = List.combine def cl in - let pr_arg ((id, _, _), arg) = - let kn = change_kn_label tpe id in - pr_projection kn ++ spc () ++ str ":=" ++ spc () ++ pr_glbexpr E1 arg - in - let args = prlist_with_sep pr_semicolon pr_arg args in - hv 0 (str "{" ++ spc () ++ args ++ spc () ++ str "}") - | (GTydDef _ | GTydOpn) -> assert false - in - hov 0 (pr_glbexpr lvl c) - - - -let pr_glbexpr c = - pr_glbexpr_gen E5 c - -(** Toplevel printers *) - -let rec subst_type subst (t : 'a glb_typexpr) = match t with -| GTypVar id -> subst.(id) -| GTypArrow (t1, t2) -> GTypArrow (subst_type subst t1, subst_type subst t2) -| GTypRef (qid, args) -> - GTypRef (qid, List.map (fun t -> subst_type subst t) args) - -let unfold kn args = - let (nparams, def) = Tac2env.interp_type kn in - match def with - | GTydDef (Some def) -> - let args = Array.of_list args in - Some (subst_type args def) - | _ -> None - -let rec kind t = match t with -| GTypVar id -> GTypVar id -| GTypRef (Other kn, tl) -> - begin match unfold kn tl with - | None -> t - | Some t -> kind t - end -| GTypArrow _ | GTypRef (Tuple _, _) -> t - -type val_printer = - { val_printer : 'a. Environ.env -> Evd.evar_map -> valexpr -> 'a glb_typexpr list -> Pp.t } - -let printers = ref KNmap.empty - -let register_val_printer kn pr = - printers := KNmap.add kn pr !printers - -open Tac2ffi - -let rec pr_valexpr env sigma v t = match kind t with -| GTypVar _ -> str "" -| GTypRef (Other kn, params) -> - let pr = try Some (KNmap.find kn !printers) with Not_found -> None in - begin match pr with - | Some pr -> pr.val_printer env sigma v params - | None -> - let n, repr = Tac2env.interp_type kn in - if KerName.equal kn t_list then - pr_val_list env sigma (to_list (fun v -> repr_to valexpr v) v) (List.hd params) - else match repr with - | GTydDef None -> str "" - | GTydDef (Some _) -> - (* Shouldn't happen thanks to kind *) - assert false - | GTydAlg alg -> - if Valexpr.is_int v then - pr_internal_constructor kn (Tac2ffi.to_int v) true - else - let (n, args) = Tac2ffi.to_block v in - let (id, tpe) = find_constructor n false alg.galg_constructors in - let knc = change_kn_label kn id in - let args = pr_constrargs env sigma params args tpe in - hv 2 (pr_constructor knc ++ spc () ++ str "(" ++ args ++ str ")") - | GTydRec rcd -> - let (_, args) = Tac2ffi.to_block v in - pr_record env sigma params args rcd - | GTydOpn -> - begin match Tac2ffi.to_open v with - | (knc, [||]) -> pr_constructor knc - | (knc, args) -> - let data = Tac2env.interp_constructor knc in - let args = pr_constrargs env sigma params args data.Tac2env.cdata_args in - hv 2 (pr_constructor knc ++ spc () ++ str "(" ++ args ++ str ")") - end - end -| GTypArrow _ -> str "" -| GTypRef (Tuple 0, []) -> str "()" -| GTypRef (Tuple _, tl) -> - let blk = Array.to_list (snd (to_block v)) in - if List.length blk == List.length tl then - let prs = List.map2 (fun v t -> pr_valexpr env sigma v t) blk tl in - hv 2 (str "(" ++ prlist_with_sep pr_comma (fun p -> p) prs ++ str ")") - else - str "" - -and pr_constrargs env sigma params args tpe = - let subst = Array.of_list params in - let tpe = List.map (fun t -> subst_type subst t) tpe in - let args = Array.to_list args in - let args = List.combine args tpe in - prlist_with_sep pr_comma (fun (v, t) -> pr_valexpr env sigma v t) args - -and pr_record env sigma params args rcd = - let subst = Array.of_list params in - let map (id, _, tpe) = (id, subst_type subst tpe) in - let rcd = List.map map rcd in - let args = Array.to_list args in - let fields = List.combine rcd args in - let pr_field ((id, t), arg) = - Id.print id ++ spc () ++ str ":=" ++ spc () ++ pr_valexpr env sigma arg t - in - str "{" ++ spc () ++ prlist_with_sep pr_semicolon pr_field fields ++ spc () ++ str "}" - -and pr_val_list env sigma args tpe = - let pr v = pr_valexpr env sigma v tpe in - str "[" ++ prlist_with_sep pr_semicolon pr args ++ str "]" - -let register_init n f = - let kn = KerName.make Tac2env.coq_prefix (Label.make n) in - register_val_printer kn { val_printer = fun env sigma v _ -> f env sigma v } - -let () = register_init "int" begin fun _ _ n -> - let n = to_int n in - Pp.int n -end - -let () = register_init "string" begin fun _ _ s -> - let s = to_string s in - Pp.quote (str (Bytes.to_string s)) -end - -let () = register_init "ident" begin fun _ _ id -> - let id = to_ident id in - str "@" ++ Id.print id -end - -let () = register_init "constr" begin fun env sigma c -> - let c = to_constr c in - let c = try Printer.pr_leconstr_env env sigma c with _ -> str "..." in - str "constr:(" ++ c ++ str ")" -end - -let () = register_init "pattern" begin fun env sigma c -> - let c = to_pattern c in - let c = try Printer.pr_lconstr_pattern_env env sigma c with _ -> str "..." in - str "pattern:(" ++ c ++ str ")" -end - -let () = register_init "message" begin fun _ _ pp -> - str "message:(" ++ to_pp pp ++ str ")" -end - -let () = register_init "err" begin fun _ _ e -> - let e = to_ext val_exn e in - let (e, _) = ExplainErr.process_vernac_interp_error ~allow_uncaught:true e in - str "err:(" ++ CErrors.print_no_report e ++ str ")" -end - -let () = - let kn = KerName.make Tac2env.coq_prefix (Label.make "array") in - let val_printer env sigma v arg = match arg with - | [arg] -> - let (_, v) = to_block v in - str "[|" ++ spc () ++ - prvect_with_sep pr_semicolon (fun a -> pr_valexpr env sigma a arg) v ++ - spc () ++ str "|]" - | _ -> assert false - in - register_val_printer kn { val_printer } diff --git a/src/tac2print.mli b/src/tac2print.mli deleted file mode 100644 index 9b9db2937d..0000000000 --- a/src/tac2print.mli +++ /dev/null @@ -1,46 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* Pp.t -val pr_glbtype_gen : ('a -> string) -> typ_level -> 'a glb_typexpr -> Pp.t -val pr_glbtype : ('a -> string) -> 'a glb_typexpr -> Pp.t - -(** {5 Printing expressions} *) - -val pr_constructor : ltac_constructor -> Pp.t -val pr_internal_constructor : type_constant -> int -> bool -> Pp.t -val pr_projection : ltac_projection -> Pp.t -val pr_glbexpr_gen : exp_level -> glb_tacexpr -> Pp.t -val pr_glbexpr : glb_tacexpr -> Pp.t - -(** {5 Printing values}*) - -type val_printer = - { val_printer : 'a. Environ.env -> Evd.evar_map -> valexpr -> 'a glb_typexpr list -> Pp.t } - -val register_val_printer : type_constant -> val_printer -> unit - -val pr_valexpr : Environ.env -> Evd.evar_map -> valexpr -> 'a glb_typexpr -> Pp.t - -(** {5 Utilities} *) - -val int_name : unit -> (int -> string) -(** Create a function that give names to integers. The names are generated on - the fly, in the order they are encountered. *) diff --git a/src/tac2qexpr.mli b/src/tac2qexpr.mli deleted file mode 100644 index 400ab1a092..0000000000 --- a/src/tac2qexpr.mli +++ /dev/null @@ -1,173 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* - CAst.make ?loc @@ CTacCst (AbsKn (Tuple 0)) -| [e] -> e -| el -> - let len = List.length el in - CAst.make ?loc @@ CTacApp (CAst.make ?loc @@ CTacCst (AbsKn (Tuple len)), el) - -let of_int {loc;v=n} = - CAst.make ?loc @@ CTacAtm (AtmInt n) - -let of_option ?loc f opt = match opt with -| None -> constructor ?loc (coq_core "None") [] -| Some e -> constructor ?loc (coq_core "Some") [f e] - -let inj_wit ?loc wit x = - CAst.make ?loc @@ CTacExt (wit, x) - -let of_variable {loc;v=id} = - let qid = Libnames.qualid_of_ident ?loc id in - if Tac2env.is_constructor qid then - CErrors.user_err ?loc (str "Invalid identifier") - else CAst.make ?loc @@ CTacRef (RelId qid) - -let of_anti f = function -| QExpr x -> f x -| QAnti id -> of_variable id - -let of_ident {loc;v=id} = inj_wit ?loc wit_ident id - -let of_constr c = - let loc = Constrexpr_ops.constr_loc c in - inj_wit ?loc wit_constr c - -let of_open_constr c = - let loc = Constrexpr_ops.constr_loc c in - inj_wit ?loc wit_open_constr c - -let of_bool ?loc b = - let c = if b then coq_core "true" else coq_core "false" in - constructor ?loc c [] - -let rec of_list ?loc f = function -| [] -> constructor (coq_core "[]") [] -| e :: l -> - constructor ?loc (coq_core "::") [f e; of_list ?loc f l] - -let of_qhyp {loc;v=h} = match h with -| QAnonHyp n -> std_constructor ?loc "AnonHyp" [of_int n] -| QNamedHyp id -> std_constructor ?loc "NamedHyp" [of_ident id] - -let of_bindings {loc;v=b} = match b with -| QNoBindings -> - std_constructor ?loc "NoBindings" [] -| QImplicitBindings tl -> - std_constructor ?loc "ImplicitBindings" [of_list ?loc of_open_constr tl] -| QExplicitBindings tl -> - let map e = of_pair (fun q -> of_anti of_qhyp q) of_open_constr e in - std_constructor ?loc "ExplicitBindings" [of_list ?loc map tl] - -let of_constr_with_bindings c = of_pair of_open_constr of_bindings c - -let rec of_intro_pattern {loc;v=pat} = match pat with -| QIntroForthcoming b -> - std_constructor ?loc "IntroForthcoming" [of_bool b] -| QIntroNaming iname -> - std_constructor ?loc "IntroNaming" [of_intro_pattern_naming iname] -| QIntroAction iact -> - std_constructor ?loc "IntroAction" [of_intro_pattern_action iact] - -and of_intro_pattern_naming {loc;v=pat} = match pat with -| QIntroIdentifier id -> - std_constructor ?loc "IntroIdentifier" [of_anti of_ident id] -| QIntroFresh id -> - std_constructor ?loc "IntroFresh" [of_anti of_ident id] -| QIntroAnonymous -> - std_constructor ?loc "IntroAnonymous" [] - -and of_intro_pattern_action {loc;v=pat} = match pat with -| QIntroWildcard -> - std_constructor ?loc "IntroWildcard" [] -| QIntroOrAndPattern pat -> - std_constructor ?loc "IntroOrAndPattern" [of_or_and_intro_pattern pat] -| QIntroInjection il -> - std_constructor ?loc "IntroInjection" [of_intro_patterns il] -| QIntroRewrite b -> - std_constructor ?loc "IntroRewrite" [of_bool ?loc b] - -and of_or_and_intro_pattern {loc;v=pat} = match pat with -| QIntroOrPattern ill -> - std_constructor ?loc "IntroOrPattern" [of_list ?loc of_intro_patterns ill] -| QIntroAndPattern il -> - std_constructor ?loc "IntroAndPattern" [of_intro_patterns il] - -and of_intro_patterns {loc;v=l} = - of_list ?loc of_intro_pattern l - -let of_hyp_location_flag ?loc = function -| Locus.InHyp -> std_constructor ?loc "InHyp" [] -| Locus.InHypTypeOnly -> std_constructor ?loc "InHypTypeOnly" [] -| Locus.InHypValueOnly -> std_constructor ?loc "InHypValueOnly" [] - -let of_occurrences {loc;v=occ} = match occ with -| QAllOccurrences -> std_constructor ?loc "AllOccurrences" [] -| QAllOccurrencesBut occs -> - let map occ = of_anti of_int occ in - let occs = of_list ?loc map occs in - std_constructor ?loc "AllOccurrencesBut" [occs] -| QNoOccurrences -> std_constructor ?loc "NoOccurrences" [] -| QOnlyOccurrences occs -> - let map occ = of_anti of_int occ in - let occs = of_list ?loc map occs in - std_constructor ?loc "OnlyOccurrences" [occs] - -let of_hyp_location ?loc ((occs, id), flag) = - of_tuple ?loc [ - of_anti of_ident id; - of_occurrences occs; - of_hyp_location_flag ?loc flag; - ] - -let of_clause {loc;v=cl} = - let hyps = of_option ?loc (fun l -> of_list ?loc of_hyp_location l) cl.q_onhyps in - let concl = of_occurrences cl.q_concl_occs in - CAst.make ?loc @@ CTacRec ([ - std_proj "on_hyps", hyps; - std_proj "on_concl", concl; - ]) - -let of_destruction_arg {loc;v=arg} = match arg with -| QElimOnConstr c -> - let arg = thunk (of_constr_with_bindings c) in - std_constructor ?loc "ElimOnConstr" [arg] -| QElimOnIdent id -> std_constructor ?loc "ElimOnIdent" [of_ident id] -| QElimOnAnonHyp n -> std_constructor ?loc "ElimOnAnonHyp" [of_int n] - -let of_induction_clause {loc;v=cl} = - let arg = of_destruction_arg cl.indcl_arg in - let eqn = of_option ?loc of_intro_pattern_naming cl.indcl_eqn in - let as_ = of_option ?loc of_or_and_intro_pattern cl.indcl_as in - let in_ = of_option ?loc of_clause cl.indcl_in in - CAst.make ?loc @@ CTacRec ([ - std_proj "indcl_arg", arg; - std_proj "indcl_eqn", eqn; - std_proj "indcl_as", as_; - std_proj "indcl_in", in_; - ]) - -let check_pattern_id ?loc id = - if Tac2env.is_constructor (Libnames.qualid_of_ident id) then - CErrors.user_err ?loc (str "Invalid pattern binding name " ++ Id.print id) - -let pattern_vars pat = - let rec aux () accu pat = match pat.CAst.v with - | Constrexpr.CPatVar id - | Constrexpr.CEvar (id, []) -> - let () = check_pattern_id ?loc:pat.CAst.loc id in - Id.Set.add id accu - | _ -> - Constrexpr_ops.fold_constr_expr_with_binders (fun _ () -> ()) aux () accu pat - in - aux () Id.Set.empty pat - -let abstract_vars loc vars tac = - let get_name = function Name id -> Some id | Anonymous -> None in - let def = try Some (List.find_map get_name vars) with Not_found -> None in - let na, tac = match def with - | None -> (Anonymous, tac) - | Some id0 -> - (* Trick: in order not to shadow a variable nor to choose an arbitrary - name, we reuse one which is going to be shadowed by the matched - variables anyways. *) - let build_bindings (n, accu) na = match na with - | Anonymous -> (n + 1, accu) - | Name _ -> - let get = global_ref ?loc (kername array_prefix "get") in - let args = [of_variable CAst.(make ?loc id0); of_int CAst.(make ?loc n)] in - let e = CAst.make ?loc @@ CTacApp (get, args) in - let accu = (CAst.make ?loc @@ CPatVar na, e) :: accu in - (n + 1, accu) - in - let (_, bnd) = List.fold_left build_bindings (0, []) vars in - let tac = CAst.make ?loc @@ CTacLet (false, bnd, tac) in - (Name id0, tac) - in - CAst.make ?loc @@ CTacFun ([CAst.make ?loc @@ CPatVar na], tac) - -let of_pattern p = - inj_wit ?loc:p.CAst.loc wit_pattern p - -let of_conversion {loc;v=c} = match c with -| QConvert c -> - let pat = of_option ?loc of_pattern None in - let c = CAst.make ?loc @@ CTacFun ([CAst.make ?loc @@ CPatVar Anonymous], of_constr c) in - of_tuple ?loc [pat; c] -| QConvertWith (pat, c) -> - let vars = pattern_vars pat in - let pat = of_option ?loc of_pattern (Some pat) in - let c = of_constr c in - (* Order is critical here *) - let vars = List.map (fun id -> Name id) (Id.Set.elements vars) in - let c = abstract_vars loc vars c in - of_tuple [pat; c] - -let of_repeat {loc;v=r} = match r with -| QPrecisely n -> std_constructor ?loc "Precisely" [of_int n] -| QUpTo n -> std_constructor ?loc "UpTo" [of_int n] -| QRepeatStar -> std_constructor ?loc "RepeatStar" [] -| QRepeatPlus -> std_constructor ?loc "RepeatPlus" [] - -let of_orient loc b = - if b then std_constructor ?loc "LTR" [] - else std_constructor ?loc "RTL" [] - -let of_rewriting {loc;v=rew} = - let orient = - let {loc;v=orient} = rew.rew_orient in - of_option ?loc (fun b -> of_orient loc b) orient - in - let repeat = of_repeat rew.rew_repeat in - let equatn = thunk (of_constr_with_bindings rew.rew_equatn) in - CAst.make ?loc @@ CTacRec ([ - std_proj "rew_orient", orient; - std_proj "rew_repeat", repeat; - std_proj "rew_equatn", equatn; - ]) - -let of_hyp ?loc id = - let hyp = global_ref ?loc (control_core "hyp") in - CAst.make ?loc @@ CTacApp (hyp, [of_ident id]) - -let of_exact_hyp ?loc id = - let refine = global_ref ?loc (control_core "refine") in - CAst.make ?loc @@ CTacApp (refine, [thunk (of_hyp ?loc id)]) - -let of_exact_var ?loc id = - let refine = global_ref ?loc (control_core "refine") in - CAst.make ?loc @@ CTacApp (refine, [thunk (of_variable id)]) - -let of_dispatch tacs = - let loc = tacs.loc in - let default = function - | Some e -> thunk e - | None -> thunk (CAst.make ?loc @@ CTacCst (AbsKn (Tuple 0))) - in - let map e = of_pair default (fun l -> of_list ?loc default l) (CAst.make ?loc e) in - of_pair (fun l -> of_list ?loc default l) (fun r -> of_option ?loc map r) tacs - -let make_red_flag l = - let open Genredexpr in - let rec add_flag red = function - | [] -> red - | {v=flag} :: lf -> - let red = match flag with - | QBeta -> { red with rBeta = true } - | QMatch -> { red with rMatch = true } - | QFix -> { red with rFix = true } - | QCofix -> { red with rCofix = true } - | QZeta -> { red with rZeta = true } - | QConst {loc;v=l} -> - if red.rDelta then - CErrors.user_err ?loc Pp.(str - "Cannot set both constants to unfold and constants not to unfold"); - { red with rConst = red.rConst @ l } - | QDeltaBut {loc;v=l} -> - if red.rConst <> [] && not red.rDelta then - CErrors.user_err ?loc Pp.(str - "Cannot set both constants to unfold and constants not to unfold"); - { red with rConst = red.rConst @ l; rDelta = true } - | QIota -> - { red with rMatch = true; rFix = true; rCofix = true } - in - add_flag red lf - in - add_flag - {rBeta = false; rMatch = false; rFix = false; rCofix = false; - rZeta = false; rDelta = false; rConst = []} - l - -let of_reference r = - let of_ref ref = - inj_wit ?loc:ref.loc wit_reference ref - in - of_anti of_ref r - -let of_strategy_flag {loc;v=flag} = - let open Genredexpr in - let flag = make_red_flag flag in - CAst.make ?loc @@ CTacRec ([ - std_proj "rBeta", of_bool ?loc flag.rBeta; - std_proj "rMatch", of_bool ?loc flag.rMatch; - std_proj "rFix", of_bool ?loc flag.rFix; - std_proj "rCofix", of_bool ?loc flag.rCofix; - std_proj "rZeta", of_bool ?loc flag.rZeta; - std_proj "rDelta", of_bool ?loc flag.rDelta; - std_proj "rConst", of_list ?loc of_reference flag.rConst; - ]) - -let of_hintdb {loc;v=hdb} = match hdb with -| QHintAll -> of_option ?loc (fun l -> of_list (fun id -> of_anti of_ident id) l) None -| QHintDbs ids -> of_option ?loc (fun l -> of_list (fun id -> of_anti of_ident id) l) (Some ids) - -let extract_name ?loc oid = match oid with -| None -> Anonymous -| Some id -> - let () = check_pattern_id ?loc id in - Name id - -(** For every branch in the matching, generate a corresponding term of type - [(match_kind * pattern * (context -> constr array -> 'a))] - where the function binds the names from the pattern to the contents of the - constr array. *) -let of_constr_matching {loc;v=m} = - let map {loc;v=({loc=ploc;v=pat}, tac)} = - let (knd, pat, na) = match pat with - | QConstrMatchPattern pat -> - let knd = constructor ?loc (pattern_core "MatchPattern") [] in - (knd, pat, Anonymous) - | QConstrMatchContext (id, pat) -> - let na = extract_name ?loc id in - let knd = constructor ?loc (pattern_core "MatchContext") [] in - (knd, pat, na) - in - let vars = pattern_vars pat in - (* Order of elements is crucial here! *) - let vars = Id.Set.elements vars in - let vars = List.map (fun id -> Name id) vars in - let e = abstract_vars loc vars tac in - let e = CAst.make ?loc @@ CTacFun ([CAst.make ?loc @@ CPatVar na], e) in - let pat = inj_wit ?loc:ploc wit_pattern pat in - of_tuple [knd; pat; e] - in - of_list ?loc map m - -(** From the patterns and the body of the branch, generate: - - a goal pattern: (constr_match list * constr_match) - - a branch function (ident array -> context array -> constr array -> context -> 'a) -*) -let of_goal_matching {loc;v=gm} = - let mk_pat {loc;v=p} = match p with - | QConstrMatchPattern pat -> - let knd = constructor ?loc (pattern_core "MatchPattern") [] in - (Anonymous, pat, knd) - | QConstrMatchContext (id, pat) -> - let na = extract_name ?loc id in - let knd = constructor ?loc (pattern_core "MatchContext") [] in - (na, pat, knd) - in - let mk_gpat {loc;v=p} = - let concl_pat = p.q_goal_match_concl in - let hyps_pats = p.q_goal_match_hyps in - let (concl_ctx, concl_pat, concl_knd) = mk_pat concl_pat in - let vars = pattern_vars concl_pat in - let map accu (na, pat) = - let (ctx, pat, knd) = mk_pat pat in - let vars = pattern_vars pat in - (Id.Set.union vars accu, (na, ctx, pat, knd)) - in - let (vars, hyps_pats) = List.fold_left_map map vars hyps_pats in - let map (_, _, pat, knd) = of_tuple [knd; of_pattern pat] in - let concl = of_tuple [concl_knd; of_pattern concl_pat] in - let r = of_tuple [of_list ?loc map hyps_pats; concl] in - let hyps = List.map (fun ({CAst.v=na}, _, _, _) -> na) hyps_pats in - let map (_, na, _, _) = na in - let hctx = List.map map hyps_pats in - (* Order of elements is crucial here! *) - let vars = Id.Set.elements vars in - let subst = List.map (fun id -> Name id) vars in - (r, hyps, hctx, subst, concl_ctx) - in - let map {loc;v=(pat, tac)} = - let (pat, hyps, hctx, subst, cctx) = mk_gpat pat in - let tac = CAst.make ?loc @@ CTacFun ([CAst.make ?loc @@ CPatVar cctx], tac) in - let tac = abstract_vars loc subst tac in - let tac = abstract_vars loc hctx tac in - let tac = abstract_vars loc hyps tac in - of_tuple ?loc [pat; tac] - in - of_list ?loc map gm - -let of_move_location {loc;v=mv} = match mv with -| QMoveAfter id -> std_constructor ?loc "MoveAfter" [of_anti of_ident id] -| QMoveBefore id -> std_constructor ?loc "MoveBefore" [of_anti of_ident id] -| QMoveFirst -> std_constructor ?loc "MoveFirst" [] -| QMoveLast -> std_constructor ?loc "MoveLast" [] - -let of_pose p = - of_pair (fun id -> of_option (fun id -> of_anti of_ident id) id) of_open_constr p - -let of_assertion {loc;v=ast} = match ast with -| QAssertType (ipat, c, tac) -> - let ipat = of_option of_intro_pattern ipat in - let c = of_constr c in - let tac = of_option thunk tac in - std_constructor ?loc "AssertType" [ipat; c; tac] -| QAssertValue (id, c) -> - let id = of_anti of_ident id in - let c = of_constr c in - std_constructor ?loc "AssertValue" [id; c] diff --git a/src/tac2quote.mli b/src/tac2quote.mli deleted file mode 100644 index 1b03dad8ec..0000000000 --- a/src/tac2quote.mli +++ /dev/null @@ -1,102 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* ltac_constructor -> raw_tacexpr list -> raw_tacexpr - -val thunk : raw_tacexpr -> raw_tacexpr - -val of_anti : ('a -> raw_tacexpr) -> 'a or_anti -> raw_tacexpr - -val of_int : int CAst.t -> raw_tacexpr - -val of_pair : ('a -> raw_tacexpr) -> ('b -> raw_tacexpr) -> ('a * 'b) CAst.t -> raw_tacexpr - -val of_tuple : ?loc:Loc.t -> raw_tacexpr list -> raw_tacexpr - -val of_variable : Id.t CAst.t -> raw_tacexpr - -val of_ident : Id.t CAst.t -> raw_tacexpr - -val of_constr : Constrexpr.constr_expr -> raw_tacexpr - -val of_open_constr : Constrexpr.constr_expr -> raw_tacexpr - -val of_list : ?loc:Loc.t -> ('a -> raw_tacexpr) -> 'a list -> raw_tacexpr - -val of_bindings : bindings -> raw_tacexpr - -val of_intro_pattern : intro_pattern -> raw_tacexpr - -val of_intro_patterns : intro_pattern list CAst.t -> raw_tacexpr - -val of_clause : clause -> raw_tacexpr - -val of_destruction_arg : destruction_arg -> raw_tacexpr - -val of_induction_clause : induction_clause -> raw_tacexpr - -val of_conversion : conversion -> raw_tacexpr - -val of_rewriting : rewriting -> raw_tacexpr - -val of_occurrences : occurrences -> raw_tacexpr - -val of_hintdb : hintdb -> raw_tacexpr - -val of_move_location : move_location -> raw_tacexpr - -val of_reference : reference or_anti -> raw_tacexpr - -val of_hyp : ?loc:Loc.t -> Id.t CAst.t -> raw_tacexpr -(** id ↦ 'Control.hyp @id' *) - -val of_exact_hyp : ?loc:Loc.t -> Id.t CAst.t -> raw_tacexpr -(** id ↦ 'Control.refine (fun () => Control.hyp @id') *) - -val of_exact_var : ?loc:Loc.t -> Id.t CAst.t -> raw_tacexpr -(** id ↦ 'Control.refine (fun () => Control.hyp @id') *) - -val of_dispatch : dispatch -> raw_tacexpr - -val of_strategy_flag : strategy_flag -> raw_tacexpr - -val of_pose : pose -> raw_tacexpr - -val of_assertion : assertion -> raw_tacexpr - -val of_constr_matching : constr_matching -> raw_tacexpr - -val of_goal_matching : goal_matching -> raw_tacexpr - -(** {5 Generic arguments} *) - -val wit_pattern : (Constrexpr.constr_expr, Pattern.constr_pattern) Arg.tag - -val wit_ident : (Id.t, Id.t) Arg.tag - -val wit_reference : (reference, GlobRef.t) Arg.tag - -val wit_constr : (Constrexpr.constr_expr, Glob_term.glob_constr) Arg.tag - -val wit_open_constr : (Constrexpr.constr_expr, Glob_term.glob_constr) Arg.tag - -val wit_ltac1 : (Ltac_plugin.Tacexpr.raw_tactic_expr, Ltac_plugin.Tacexpr.glob_tactic_expr) Arg.tag -(** Ltac1 AST quotation, seen as a 'tactic'. Its type is unit in Ltac2. *) - -val wit_ltac1val : (Ltac_plugin.Tacexpr.raw_tactic_expr, Ltac_plugin.Tacexpr.glob_tactic_expr) Arg.tag -(** Ltac1 AST quotation, seen as a value-returning expression, with type Ltac1.t. *) diff --git a/src/tac2stdlib.ml b/src/tac2stdlib.ml deleted file mode 100644 index ffef2c05fd..0000000000 --- a/src/tac2stdlib.ml +++ /dev/null @@ -1,578 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* assert false) f - -let return x = Proofview.tclUNIT x -let v_unit = Value.of_unit () -let thaw r f = Tac2ffi.app_fun1 f unit r () -let uthaw r f = Tac2ffi.app_fun1 (to_fun1 unit r f) unit r () -let thunk r = fun1 unit r - -let to_name c = match Value.to_option Value.to_ident c with -| None -> Anonymous -| Some id -> Name id - -let name = make_to_repr to_name - -let to_occurrences = function -| ValInt 0 -> AllOccurrences -| ValBlk (0, [| vl |]) -> AllOccurrencesBut (Value.to_list Value.to_int vl) -| ValInt 1 -> NoOccurrences -| ValBlk (1, [| vl |]) -> OnlyOccurrences (Value.to_list Value.to_int vl) -| _ -> assert false - -let occurrences = make_to_repr to_occurrences - -let to_hyp_location_flag v = match Value.to_int v with -| 0 -> InHyp -| 1 -> InHypTypeOnly -| 2 -> InHypValueOnly -| _ -> assert false - -let to_clause v = match Value.to_tuple v with -| [| hyps; concl |] -> - let cast v = match Value.to_tuple v with - | [| hyp; occ; flag |] -> - (Value.to_ident hyp, to_occurrences occ, to_hyp_location_flag flag) - | _ -> assert false - in - let hyps = Value.to_option (fun h -> Value.to_list cast h) hyps in - { onhyps = hyps; concl_occs = to_occurrences concl; } -| _ -> assert false - -let clause = make_to_repr to_clause - -let to_red_flag v = match Value.to_tuple v with -| [| beta; iota; fix; cofix; zeta; delta; const |] -> - { - rBeta = Value.to_bool beta; - rMatch = Value.to_bool iota; - rFix = Value.to_bool fix; - rCofix = Value.to_bool cofix; - rZeta = Value.to_bool zeta; - rDelta = Value.to_bool delta; - rConst = Value.to_list Value.to_reference const; - } -| _ -> assert false - -let red_flags = make_to_repr to_red_flag - -let pattern_with_occs = pair pattern occurrences - -let constr_with_occs = pair constr occurrences - -let reference_with_occs = pair reference occurrences - -let rec to_intro_pattern v = match Value.to_block v with -| (0, [| b |]) -> IntroForthcoming (Value.to_bool b) -| (1, [| pat |]) -> IntroNaming (to_intro_pattern_naming pat) -| (2, [| act |]) -> IntroAction (to_intro_pattern_action act) -| _ -> assert false - -and to_intro_pattern_naming = function -| ValBlk (0, [| id |]) -> IntroIdentifier (Value.to_ident id) -| ValBlk (1, [| id |]) -> IntroFresh (Value.to_ident id) -| ValInt 0 -> IntroAnonymous -| _ -> assert false - -and to_intro_pattern_action = function -| ValInt 0 -> IntroWildcard -| ValBlk (0, [| op |]) -> IntroOrAndPattern (to_or_and_intro_pattern op) -| ValBlk (1, [| inj |]) -> - let map ipat = to_intro_pattern ipat in - IntroInjection (Value.to_list map inj) -| ValBlk (2, [| c; ipat |]) -> - let c = Value.to_fun1 Value.unit Value.constr c in - IntroApplyOn (c, to_intro_pattern ipat) -| ValBlk (3, [| b |]) -> IntroRewrite (Value.to_bool b) -| _ -> assert false - -and to_or_and_intro_pattern v = match Value.to_block v with -| (0, [| ill |]) -> - IntroOrPattern (Value.to_list to_intro_patterns ill) -| (1, [| il |]) -> - IntroAndPattern (to_intro_patterns il) -| _ -> assert false - -and to_intro_patterns il = - Value.to_list to_intro_pattern il - -let intro_pattern = make_to_repr to_intro_pattern - -let intro_patterns = make_to_repr to_intro_patterns - -let to_destruction_arg v = match Value.to_block v with -| (0, [| c |]) -> - let c = uthaw constr_with_bindings c in - ElimOnConstr c -| (1, [| id |]) -> ElimOnIdent (Value.to_ident id) -| (2, [| n |]) -> ElimOnAnonHyp (Value.to_int n) -| _ -> assert false - -let destruction_arg = make_to_repr to_destruction_arg - -let to_induction_clause v = match Value.to_tuple v with -| [| arg; eqn; as_; in_ |] -> - let arg = to_destruction_arg arg in - let eqn = Value.to_option to_intro_pattern_naming eqn in - let as_ = Value.to_option to_or_and_intro_pattern as_ in - let in_ = Value.to_option to_clause in_ in - (arg, eqn, as_, in_) -| _ -> - assert false - -let induction_clause = make_to_repr to_induction_clause - -let to_assertion v = match Value.to_block v with -| (0, [| ipat; t; tac |]) -> - let to_tac t = Value.to_fun1 Value.unit Value.unit t in - let ipat = Value.to_option to_intro_pattern ipat in - let t = Value.to_constr t in - let tac = Value.to_option to_tac tac in - AssertType (ipat, t, tac) -| (1, [| id; c |]) -> - AssertValue (Value.to_ident id, Value.to_constr c) -| _ -> assert false - -let assertion = make_to_repr to_assertion - -let to_multi = function -| ValBlk (0, [| n |]) -> Precisely (Value.to_int n) -| ValBlk (1, [| n |]) -> UpTo (Value.to_int n) -| ValInt 0 -> RepeatStar -| ValInt 1 -> RepeatPlus -| _ -> assert false - -let to_rewriting v = match Value.to_tuple v with -| [| orient; repeat; c |] -> - let orient = Value.to_option Value.to_bool orient in - let repeat = to_multi repeat in - let c = uthaw constr_with_bindings c in - (orient, repeat, c) -| _ -> assert false - -let rewriting = make_to_repr to_rewriting - -let to_debug v = match Value.to_int v with -| 0 -> Hints.Off -| 1 -> Hints.Info -| 2 -> Hints.Debug -| _ -> assert false - -let debug = make_to_repr to_debug - -let to_strategy v = match Value.to_int v with -| 0 -> Class_tactics.Bfs -| 1 -> Class_tactics.Dfs -| _ -> assert false - -let strategy = make_to_repr to_strategy - -let to_inversion_kind v = match Value.to_int v with -| 0 -> Inv.SimpleInversion -| 1 -> Inv.FullInversion -| 2 -> Inv.FullInversionClear -| _ -> assert false - -let inversion_kind = make_to_repr to_inversion_kind - -let to_move_location = function -| ValInt 0 -> Logic.MoveFirst -| ValInt 1 -> Logic.MoveLast -| ValBlk (0, [|id|]) -> Logic.MoveAfter (Value.to_ident id) -| ValBlk (1, [|id|]) -> Logic.MoveBefore (Value.to_ident id) -| _ -> assert false - -let move_location = make_to_repr to_move_location - -let to_generalize_arg v = match Value.to_tuple v with -| [| c; occs; na |] -> - (Value.to_constr c, to_occurrences occs, to_name na) -| _ -> assert false - -let generalize_arg = make_to_repr to_generalize_arg - -(** Standard tactics sharing their implementation with Ltac1 *) - -let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } - -let lift tac = tac <*> return v_unit - -let define_prim0 name tac = - let tac _ = lift tac in - Tac2env.define_primitive (pname name) (mk_closure arity_one tac) - -let define_prim1 name r0 f = - let tac x = lift (f (Value.repr_to r0 x)) in - Tac2env.define_primitive (pname name) (mk_closure arity_one tac) - -let define_prim2 name r0 r1 f = - let tac x y = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y)) in - Tac2env.define_primitive (pname name) (mk_closure (arity_suc arity_one) tac) - -let define_prim3 name r0 r1 r2 f = - let tac x y z = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y) (Value.repr_to r2 z)) in - Tac2env.define_primitive (pname name) (mk_closure (arity_suc (arity_suc arity_one)) tac) - -let define_prim4 name r0 r1 r2 r3 f = - let tac x y z u = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y) (Value.repr_to r2 z) (Value.repr_to r3 u)) in - Tac2env.define_primitive (pname name) (mk_closure (arity_suc (arity_suc (arity_suc arity_one))) tac) - -let define_prim5 name r0 r1 r2 r3 r4 f = - let tac x y z u v = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y) (Value.repr_to r2 z) (Value.repr_to r3 u) (Value.repr_to r4 v)) in - Tac2env.define_primitive (pname name) (mk_closure (arity_suc (arity_suc (arity_suc (arity_suc arity_one)))) tac) - -(** Tactics from Tacexpr *) - -let () = define_prim2 "tac_intros" bool intro_patterns begin fun ev ipat -> - Tac2tactics.intros_patterns ev ipat -end - -let () = define_prim4 "tac_apply" bool bool (list (thunk constr_with_bindings)) (option (pair ident (option intro_pattern))) begin fun adv ev cb ipat -> - Tac2tactics.apply adv ev cb ipat -end - -let () = define_prim3 "tac_elim" bool constr_with_bindings (option constr_with_bindings) begin fun ev c copt -> - Tac2tactics.elim ev c copt -end - -let () = define_prim2 "tac_case" bool constr_with_bindings begin fun ev c -> - Tac2tactics.general_case_analysis ev c -end - -let () = define_prim1 "tac_generalize" (list generalize_arg) begin fun cl -> - Tac2tactics.generalize cl -end - -let () = define_prim1 "tac_assert" assertion begin fun ast -> - Tac2tactics.assert_ ast -end - -let () = define_prim3 "tac_enough" constr (option (option (thunk unit))) (option intro_pattern) begin fun c tac ipat -> - let tac = Option.map (fun o -> Option.map (fun f -> thaw unit f) o) tac in - Tac2tactics.forward false tac ipat c -end - -let () = define_prim2 "tac_pose" name constr begin fun na c -> - Tactics.letin_tac None na c None Locusops.nowhere -end - -let () = define_prim3 "tac_set" bool (thunk (pair name constr)) clause begin fun ev p cl -> - Proofview.tclEVARMAP >>= fun sigma -> - thaw (pair name constr) p >>= fun (na, c) -> - Tac2tactics.letin_pat_tac ev None na (sigma, c) cl -end - -let () = define_prim5 "tac_remember" bool name (thunk constr) (option intro_pattern) clause begin fun ev na c eqpat cl -> - let eqpat = Option.default (IntroNaming IntroAnonymous) eqpat in - match eqpat with - | IntroNaming eqpat -> - Proofview.tclEVARMAP >>= fun sigma -> - thaw constr c >>= fun c -> - Tac2tactics.letin_pat_tac ev (Some (true, eqpat)) na (sigma, c) cl - | _ -> - Tacticals.New.tclZEROMSG (Pp.str "Invalid pattern for remember") -end - -let () = define_prim3 "tac_destruct" bool (list induction_clause) (option constr_with_bindings) begin fun ev ic using -> - Tac2tactics.induction_destruct false ev ic using -end - -let () = define_prim3 "tac_induction" bool (list induction_clause) (option constr_with_bindings) begin fun ev ic using -> - Tac2tactics.induction_destruct true ev ic using -end - -let () = define_prim1 "tac_red" clause begin fun cl -> - Tac2tactics.reduce (Red false) cl -end - -let () = define_prim1 "tac_hnf" clause begin fun cl -> - Tac2tactics.reduce Hnf cl -end - -let () = define_prim3 "tac_simpl" red_flags (option pattern_with_occs) clause begin fun flags where cl -> - Tac2tactics.simpl flags where cl -end - -let () = define_prim2 "tac_cbv" red_flags clause begin fun flags cl -> - Tac2tactics.cbv flags cl -end - -let () = define_prim2 "tac_cbn" red_flags clause begin fun flags cl -> - Tac2tactics.cbn flags cl -end - -let () = define_prim2 "tac_lazy" red_flags clause begin fun flags cl -> - Tac2tactics.lazy_ flags cl -end - -let () = define_prim2 "tac_unfold" (list reference_with_occs) clause begin fun refs cl -> - Tac2tactics.unfold refs cl -end - -let () = define_prim2 "tac_fold" (list constr) clause begin fun args cl -> - Tac2tactics.reduce (Fold args) cl -end - -let () = define_prim2 "tac_pattern" (list constr_with_occs) clause begin fun where cl -> - Tac2tactics.pattern where cl -end - -let () = define_prim2 "tac_vm" (option pattern_with_occs) clause begin fun where cl -> - Tac2tactics.vm where cl -end - -let () = define_prim2 "tac_native" (option pattern_with_occs) clause begin fun where cl -> - Tac2tactics.native where cl -end - -(** Reduction functions *) - -let lift tac = tac >>= fun c -> Proofview.tclUNIT (Value.of_constr c) - -let define_red1 name r0 f = - let tac x = lift (f (Value.repr_to r0 x)) in - Tac2env.define_primitive (pname name) (mk_closure arity_one tac) - -let define_red2 name r0 r1 f = - let tac x y = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y)) in - Tac2env.define_primitive (pname name) (mk_closure (arity_suc arity_one) tac) - -let define_red3 name r0 r1 r2 f = - let tac x y z = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y) (Value.repr_to r2 z)) in - Tac2env.define_primitive (pname name) (mk_closure (arity_suc (arity_suc arity_one)) tac) - -let () = define_red1 "eval_red" constr begin fun c -> - Tac2tactics.eval_red c -end - -let () = define_red1 "eval_hnf" constr begin fun c -> - Tac2tactics.eval_hnf c -end - -let () = define_red3 "eval_simpl" red_flags (option pattern_with_occs) constr begin fun flags where c -> - Tac2tactics.eval_simpl flags where c -end - -let () = define_red2 "eval_cbv" red_flags constr begin fun flags c -> - Tac2tactics.eval_cbv flags c -end - -let () = define_red2 "eval_cbn" red_flags constr begin fun flags c -> - Tac2tactics.eval_cbn flags c -end - -let () = define_red2 "eval_lazy" red_flags constr begin fun flags c -> - Tac2tactics.eval_lazy flags c -end - -let () = define_red2 "eval_unfold" (list reference_with_occs) constr begin fun refs c -> - Tac2tactics.eval_unfold refs c -end - -let () = define_red2 "eval_fold" (list constr) constr begin fun args c -> - Tac2tactics.eval_fold args c -end - -let () = define_red2 "eval_pattern" (list constr_with_occs) constr begin fun where c -> - Tac2tactics.eval_pattern where c -end - -let () = define_red2 "eval_vm" (option pattern_with_occs) constr begin fun where c -> - Tac2tactics.eval_vm where c -end - -let () = define_red2 "eval_native" (option pattern_with_occs) constr begin fun where c -> - Tac2tactics.eval_native where c -end - -let () = define_prim3 "tac_change" (option pattern) (fun1 (array constr) constr) clause begin fun pat c cl -> - Tac2tactics.change pat c cl -end - -let () = define_prim4 "tac_rewrite" bool (list rewriting) clause (option (thunk unit)) begin fun ev rw cl by -> - Tac2tactics.rewrite ev rw cl by -end - -let () = define_prim4 "tac_inversion" inversion_kind destruction_arg (option intro_pattern) (option (list ident)) begin fun knd arg pat ids -> - Tac2tactics.inversion knd arg pat ids -end - -(** Tactics from coretactics *) - -let () = define_prim0 "tac_reflexivity" Tactics.intros_reflexivity - -let () = define_prim2 "tac_move" ident move_location begin fun id mv -> - Tactics.move_hyp id mv -end - -let () = define_prim2 "tac_intro" (option ident) (option move_location) begin fun id mv -> - let mv = Option.default Logic.MoveLast mv in - Tactics.intro_move id mv -end - -(* - -TACTIC EXTEND exact - [ "exact" casted_constr(c) ] -> [ Tactics.exact_no_check c ] -END - -*) - -let () = define_prim0 "tac_assumption" Tactics.assumption - -let () = define_prim1 "tac_transitivity" constr begin fun c -> - Tactics.intros_transitivity (Some c) -end - -let () = define_prim0 "tac_etransitivity" (Tactics.intros_transitivity None) - -let () = define_prim1 "tac_cut" constr begin fun c -> - Tactics.cut c -end - -let () = define_prim2 "tac_left" bool bindings begin fun ev bnd -> - Tac2tactics.left_with_bindings ev bnd -end -let () = define_prim2 "tac_right" bool bindings begin fun ev bnd -> - Tac2tactics.right_with_bindings ev bnd -end - -let () = define_prim1 "tac_introsuntil" qhyp begin fun h -> - Tactics.intros_until h -end - -let () = define_prim1 "tac_exactnocheck" constr begin fun c -> - Tactics.exact_no_check c -end - -let () = define_prim1 "tac_vmcastnocheck" constr begin fun c -> - Tactics.vm_cast_no_check c -end - -let () = define_prim1 "tac_nativecastnocheck" constr begin fun c -> - Tactics.native_cast_no_check c -end - -let () = define_prim1 "tac_constructor" bool begin fun ev -> - Tactics.any_constructor ev None -end - -let () = define_prim3 "tac_constructorn" bool int bindings begin fun ev n bnd -> - Tac2tactics.constructor_tac ev None n bnd -end - -let () = define_prim2 "tac_specialize" constr_with_bindings (option intro_pattern) begin fun c ipat -> - Tac2tactics.specialize c ipat -end - -let () = define_prim1 "tac_symmetry" clause begin fun cl -> - Tac2tactics.symmetry cl -end - -let () = define_prim2 "tac_split" bool bindings begin fun ev bnd -> - Tac2tactics.split_with_bindings ev bnd -end - -let () = define_prim1 "tac_rename" (list (pair ident ident)) begin fun ids -> - Tactics.rename_hyp ids -end - -let () = define_prim1 "tac_revert" (list ident) begin fun ids -> - Tactics.revert ids -end - -let () = define_prim0 "tac_admit" Proofview.give_up - -let () = define_prim2 "tac_fix" ident int begin fun ident n -> - Tactics.fix ident n -end - -let () = define_prim1 "tac_cofix" ident begin fun ident -> - Tactics.cofix ident -end - -let () = define_prim1 "tac_clear" (list ident) begin fun ids -> - Tactics.clear ids -end - -let () = define_prim1 "tac_keep" (list ident) begin fun ids -> - Tactics.keep ids -end - -let () = define_prim1 "tac_clearbody" (list ident) begin fun ids -> - Tactics.clear_body ids -end - -(** Tactics from extratactics *) - -let () = define_prim2 "tac_discriminate" bool (option destruction_arg) begin fun ev arg -> - Tac2tactics.discriminate ev arg -end - -let () = define_prim3 "tac_injection" bool (option intro_patterns) (option destruction_arg) begin fun ev ipat arg -> - Tac2tactics.injection ev ipat arg -end - -let () = define_prim1 "tac_absurd" constr begin fun c -> - Contradiction.absurd c -end - -let () = define_prim1 "tac_contradiction" (option constr_with_bindings) begin fun c -> - Tac2tactics.contradiction c -end - -let () = define_prim4 "tac_autorewrite" bool (option (thunk unit)) (list ident) clause begin fun all by ids cl -> - Tac2tactics.autorewrite ~all by ids cl -end - -let () = define_prim1 "tac_subst" (list ident) begin fun ids -> - Equality.subst ids -end - -let () = define_prim0 "tac_substall" (return () >>= fun () -> Equality.subst_all ()) - -(** Auto *) - -let () = define_prim3 "tac_trivial" debug (list (thunk constr)) (option (list ident)) begin fun dbg lems dbs -> - Tac2tactics.trivial dbg lems dbs -end - -let () = define_prim5 "tac_eauto" debug (option int) (option int) (list (thunk constr)) (option (list ident)) begin fun dbg n p lems dbs -> - Tac2tactics.eauto dbg n p lems dbs -end - -let () = define_prim4 "tac_auto" debug (option int) (list (thunk constr)) (option (list ident)) begin fun dbg n lems dbs -> - Tac2tactics.auto dbg n lems dbs -end - -let () = define_prim4 "tac_newauto" debug (option int) (list (thunk constr)) (option (list ident)) begin fun dbg n lems dbs -> - Tac2tactics.new_auto dbg n lems dbs -end - -let () = define_prim3 "tac_typeclasses_eauto" (option strategy) (option int) (option (list ident)) begin fun str n dbs -> - Tac2tactics.typeclasses_eauto str n dbs -end - -(** Firstorder *) - -let () = define_prim3 "tac_firstorder" (option (thunk unit)) (list reference) (list ident) begin fun tac refs ids -> - Tac2tactics.firstorder tac refs ids -end diff --git a/src/tac2stdlib.mli b/src/tac2stdlib.mli deleted file mode 100644 index 927b57074d..0000000000 --- a/src/tac2stdlib.mli +++ /dev/null @@ -1,9 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* Tactypes.ImplicitBindings l -| ExplicitBindings l -> - let l = List.map CAst.make l in - Tactypes.ExplicitBindings l -| NoBindings -> Tactypes.NoBindings - -let mk_with_bindings (x, b) = (x, mk_bindings b) - -let rec mk_intro_pattern = function -| IntroForthcoming b -> CAst.make @@ Tactypes.IntroForthcoming b -| IntroNaming ipat -> CAst.make @@ Tactypes.IntroNaming (mk_intro_pattern_naming ipat) -| IntroAction ipat -> CAst.make @@ Tactypes.IntroAction (mk_intro_pattern_action ipat) - -and mk_intro_pattern_naming = function -| IntroIdentifier id -> Namegen.IntroIdentifier id -| IntroFresh id -> Namegen.IntroFresh id -| IntroAnonymous -> Namegen.IntroAnonymous - -and mk_intro_pattern_action = function -| IntroWildcard -> Tactypes.IntroWildcard -| IntroOrAndPattern ipat -> Tactypes.IntroOrAndPattern (mk_or_and_intro_pattern ipat) -| IntroInjection ipats -> Tactypes.IntroInjection (List.map mk_intro_pattern ipats) -| IntroApplyOn (c, ipat) -> - let c = CAst.make @@ delayed_of_thunk Tac2ffi.constr c in - Tactypes.IntroApplyOn (c, mk_intro_pattern ipat) -| IntroRewrite b -> Tactypes.IntroRewrite b - -and mk_or_and_intro_pattern = function -| IntroOrPattern ipatss -> - Tactypes.IntroOrPattern (List.map (fun ipat -> List.map mk_intro_pattern ipat) ipatss) -| IntroAndPattern ipats -> - Tactypes.IntroAndPattern (List.map mk_intro_pattern ipats) - -let mk_intro_patterns ipat = List.map mk_intro_pattern ipat - -let mk_occurrences f = function -| AllOccurrences -> Locus.AllOccurrences -| AllOccurrencesBut l -> Locus.AllOccurrencesBut (List.map f l) -| NoOccurrences -> Locus.NoOccurrences -| OnlyOccurrences l -> Locus.OnlyOccurrences (List.map f l) - -let mk_occurrences_expr occ = - mk_occurrences (fun i -> Locus.ArgArg i) occ - -let mk_hyp_location (id, occs, h) = - ((mk_occurrences_expr occs, id), h) - -let mk_clause cl = { - Locus.onhyps = Option.map (fun l -> List.map mk_hyp_location l) cl.onhyps; - Locus.concl_occs = mk_occurrences_expr cl.concl_occs; -} - -let intros_patterns ev ipat = - let ipat = mk_intro_patterns ipat in - Tactics.intros_patterns ev ipat - -let apply adv ev cb cl = - let map c = - let c = thaw constr_with_bindings c >>= fun p -> return (mk_with_bindings p) in - None, CAst.make (delayed_of_tactic c) - in - let cb = List.map map cb in - match cl with - | None -> Tactics.apply_with_delayed_bindings_gen adv ev cb - | Some (id, cl) -> - let cl = Option.map mk_intro_pattern cl in - Tactics.apply_delayed_in adv ev id cb cl - -let mk_destruction_arg = function -| ElimOnConstr c -> - let c = c >>= fun c -> return (mk_with_bindings c) in - Tactics.ElimOnConstr (delayed_of_tactic c) -| ElimOnIdent id -> Tactics.ElimOnIdent CAst.(make id) -| ElimOnAnonHyp n -> Tactics.ElimOnAnonHyp n - -let mk_induction_clause (arg, eqn, as_, occ) = - let eqn = Option.map (fun ipat -> CAst.make @@ mk_intro_pattern_naming ipat) eqn in - let as_ = Option.map (fun ipat -> CAst.make @@ mk_or_and_intro_pattern ipat) as_ in - let occ = Option.map mk_clause occ in - ((None, mk_destruction_arg arg), (eqn, as_), occ) - -let induction_destruct isrec ev (ic : induction_clause list) using = - let ic = List.map mk_induction_clause ic in - let using = Option.map mk_with_bindings using in - Tactics.induction_destruct isrec ev (ic, using) - -let elim ev c copt = - let c = mk_with_bindings c in - let copt = Option.map mk_with_bindings copt in - Tactics.elim ev None c copt - -let generalize pl = - let mk_occ occs = mk_occurrences (fun i -> i) occs in - let pl = List.map (fun (c, occs, na) -> (mk_occ occs, c), na) pl in - Tactics.new_generalize_gen pl - -let general_case_analysis ev c = - let c = mk_with_bindings c in - Tactics.general_case_analysis ev None c - -let constructor_tac ev n i bnd = - let bnd = mk_bindings bnd in - Tactics.constructor_tac ev n i bnd - -let left_with_bindings ev bnd = - let bnd = mk_bindings bnd in - Tactics.left_with_bindings ev bnd - -let right_with_bindings ev bnd = - let bnd = mk_bindings bnd in - Tactics.right_with_bindings ev bnd - -let split_with_bindings ev bnd = - let bnd = mk_bindings bnd in - Tactics.split_with_bindings ev [bnd] - -let specialize c pat = - let c = mk_with_bindings c in - let pat = Option.map mk_intro_pattern pat in - Tactics.specialize c pat - -let change pat c cl = - let open Tac2ffi in - Proofview.Goal.enter begin fun gl -> - let c subst env sigma = - let subst = Array.map_of_list snd (Id.Map.bindings subst) in - delayed_of_tactic (Tac2ffi.app_fun1 c (array constr) constr subst) env sigma - in - let cl = mk_clause cl in - Tactics.change pat c cl - end - -let rewrite ev rw cl by = - let map_rw (orient, repeat, c) = - let c = c >>= fun c -> return (mk_with_bindings c) in - (Option.default true orient, repeat, None, delayed_of_tactic c) - in - let rw = List.map map_rw rw in - let cl = mk_clause cl in - let by = Option.map (fun tac -> Tacticals.New.tclCOMPLETE (thaw Tac2ffi.unit tac), Equality.Naive) by in - Equality.general_multi_rewrite ev rw cl by - -let symmetry cl = - let cl = mk_clause cl in - Tactics.intros_symmetry cl - -let forward fst tac ipat c = - let ipat = Option.map mk_intro_pattern ipat in - Tactics.forward fst tac ipat c - -let assert_ = function -| AssertValue (id, c) -> - let ipat = CAst.make @@ Tactypes.IntroNaming (Namegen.IntroIdentifier id) in - Tactics.forward true None (Some ipat) c -| AssertType (ipat, c, tac) -> - let ipat = Option.map mk_intro_pattern ipat in - let tac = Option.map (fun tac -> thaw Tac2ffi.unit tac) tac in - Tactics.forward true (Some tac) ipat c - -let letin_pat_tac ev ipat na c cl = - let ipat = Option.map (fun (b, ipat) -> (b, CAst.make @@ mk_intro_pattern_naming ipat)) ipat in - let cl = mk_clause cl in - Tactics.letin_pat_tac ev ipat na c cl - -(** Ltac interface treats differently global references than other term - arguments in reduction expressions. In Ltac1, this is done at parsing time. - Instead, we parse indifferently any pattern and dispatch when the tactic is - called. *) -let map_pattern_with_occs (pat, occ) = match pat with -| Pattern.PRef (ConstRef cst) -> (mk_occurrences_expr occ, Inl (EvalConstRef cst)) -| Pattern.PRef (VarRef id) -> (mk_occurrences_expr occ, Inl (EvalVarRef id)) -| _ -> (mk_occurrences_expr occ, Inr pat) - -let get_evaluable_reference = function -| VarRef id -> Proofview.tclUNIT (EvalVarRef id) -| ConstRef cst -> Proofview.tclUNIT (EvalConstRef cst) -| r -> - Tacticals.New.tclZEROMSG (str "Cannot coerce" ++ spc () ++ - Nametab.pr_global_env Id.Set.empty r ++ spc () ++ - str "to an evaluable reference.") - -let reduce r cl = - let cl = mk_clause cl in - Tactics.reduce r cl - -let simpl flags where cl = - let where = Option.map map_pattern_with_occs where in - let cl = mk_clause cl in - Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> - let flags = { flags with rConst } in - Tactics.reduce (Simpl (flags, where)) cl - -let cbv flags cl = - let cl = mk_clause cl in - Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> - let flags = { flags with rConst } in - Tactics.reduce (Cbv flags) cl - -let cbn flags cl = - let cl = mk_clause cl in - Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> - let flags = { flags with rConst } in - Tactics.reduce (Cbn flags) cl - -let lazy_ flags cl = - let cl = mk_clause cl in - Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> - let flags = { flags with rConst } in - Tactics.reduce (Lazy flags) cl - -let unfold occs cl = - let cl = mk_clause cl in - let map (gr, occ) = - let occ = mk_occurrences_expr occ in - get_evaluable_reference gr >>= fun gr -> Proofview.tclUNIT (occ, gr) - in - Proofview.Monad.List.map map occs >>= fun occs -> - Tactics.reduce (Unfold occs) cl - -let pattern where cl = - let where = List.map (fun (c, occ) -> (mk_occurrences_expr occ, c)) where in - let cl = mk_clause cl in - Tactics.reduce (Pattern where) cl - -let vm where cl = - let where = Option.map map_pattern_with_occs where in - let cl = mk_clause cl in - Tactics.reduce (CbvVm where) cl - -let native where cl = - let where = Option.map map_pattern_with_occs where in - let cl = mk_clause cl in - Tactics.reduce (CbvNative where) cl - -let eval_fun red c = - Tac2core.pf_apply begin fun env sigma -> - let (redfun, _) = Redexpr.reduction_of_red_expr env red in - let (sigma, ans) = redfun env sigma c in - Proofview.Unsafe.tclEVARS sigma >>= fun () -> - Proofview.tclUNIT ans - end - -let eval_red c = - eval_fun (Red false) c - -let eval_hnf c = - eval_fun Hnf c - -let eval_simpl flags where c = - let where = Option.map map_pattern_with_occs where in - Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> - let flags = { flags with rConst } in - eval_fun (Simpl (flags, where)) c - -let eval_cbv flags c = - Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> - let flags = { flags with rConst } in - eval_fun (Cbv flags) c - -let eval_cbn flags c = - Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> - let flags = { flags with rConst } in - eval_fun (Cbn flags) c - -let eval_lazy flags c = - Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> - let flags = { flags with rConst } in - eval_fun (Lazy flags) c - -let eval_unfold occs c = - let map (gr, occ) = - let occ = mk_occurrences_expr occ in - get_evaluable_reference gr >>= fun gr -> Proofview.tclUNIT (occ, gr) - in - Proofview.Monad.List.map map occs >>= fun occs -> - eval_fun (Unfold occs) c - -let eval_fold cl c = - eval_fun (Fold cl) c - -let eval_pattern where c = - let where = List.map (fun (pat, occ) -> (mk_occurrences_expr occ, pat)) where in - eval_fun (Pattern where) c - -let eval_vm where c = - let where = Option.map map_pattern_with_occs where in - eval_fun (CbvVm where) c - -let eval_native where c = - let where = Option.map map_pattern_with_occs where in - eval_fun (CbvNative where) c - -let on_destruction_arg tac ev arg = - Proofview.Goal.enter begin fun gl -> - match arg with - | None -> tac ev None - | Some (clear, arg) -> - let arg = match arg with - | ElimOnConstr c -> - let env = Proofview.Goal.env gl in - Proofview.tclEVARMAP >>= fun sigma -> - c >>= fun (c, lbind) -> - let lbind = mk_bindings lbind in - Proofview.tclEVARMAP >>= fun sigma' -> - let flags = tactic_infer_flags ev in - let (sigma', c) = Unification.finish_evar_resolution ~flags env sigma' (sigma, c) in - Proofview.tclUNIT (Some sigma', Tactics.ElimOnConstr (c, lbind)) - | ElimOnIdent id -> Proofview.tclUNIT (None, Tactics.ElimOnIdent CAst.(make id)) - | ElimOnAnonHyp n -> Proofview.tclUNIT (None, Tactics.ElimOnAnonHyp n) - in - arg >>= fun (sigma', arg) -> - let arg = Some (clear, arg) in - match sigma' with - | None -> tac ev arg - | Some sigma' -> - Tacticals.New.tclWITHHOLES ev (tac ev arg) sigma' - end - -let discriminate ev arg = - let arg = Option.map (fun arg -> None, arg) arg in - on_destruction_arg Equality.discr_tac ev arg - -let injection ev ipat arg = - let arg = Option.map (fun arg -> None, arg) arg in - let ipat = Option.map mk_intro_patterns ipat in - let tac ev arg = Equality.injClause None ipat ev arg in - on_destruction_arg tac ev arg - -let autorewrite ~all by ids cl = - let conds = if all then Some Equality.AllMatches else None in - let ids = List.map Id.to_string ids in - let cl = mk_clause cl in - match by with - | None -> Autorewrite.auto_multi_rewrite ?conds ids cl - | Some by -> - let by = thaw Tac2ffi.unit by in - Autorewrite.auto_multi_rewrite_with ?conds by ids cl - -(** Auto *) - -let trivial debug lems dbs = - let lems = List.map (fun c -> delayed_of_thunk Tac2ffi.constr c) lems in - let dbs = Option.map (fun l -> List.map Id.to_string l) dbs in - Auto.h_trivial ~debug lems dbs - -let auto debug n lems dbs = - let lems = List.map (fun c -> delayed_of_thunk Tac2ffi.constr c) lems in - let dbs = Option.map (fun l -> List.map Id.to_string l) dbs in - Auto.h_auto ~debug n lems dbs - -let new_auto debug n lems dbs = - let make_depth n = snd (Eauto.make_dimension n None) in - let lems = List.map (fun c -> delayed_of_thunk Tac2ffi.constr c) lems in - match dbs with - | None -> Auto.new_full_auto ~debug (make_depth n) lems - | Some dbs -> - let dbs = List.map Id.to_string dbs in - Auto.new_auto ~debug (make_depth n) lems dbs - -let eauto debug n p lems dbs = - let lems = List.map (fun c -> delayed_of_thunk Tac2ffi.constr c) lems in - let dbs = Option.map (fun l -> List.map Id.to_string l) dbs in - Eauto.gen_eauto (Eauto.make_dimension n p) lems dbs - -let typeclasses_eauto strategy depth dbs = - let only_classes, dbs = match dbs with - | None -> - true, [Class_tactics.typeclasses_db] - | Some dbs -> - let dbs = List.map Id.to_string dbs in - false, dbs - in - Class_tactics.typeclasses_eauto ~only_classes ?strategy ~depth dbs - -(** Inversion *) - -let inversion knd arg pat ids = - let ids = match ids with - | None -> [] - | Some l -> l - in - begin match pat with - | None -> Proofview.tclUNIT None - | Some (IntroAction (IntroOrAndPattern p)) -> - Proofview.tclUNIT (Some (CAst.make @@ mk_or_and_intro_pattern p)) - | Some _ -> - Tacticals.New.tclZEROMSG (str "Inversion only accept disjunctive patterns") - end >>= fun pat -> - let inversion _ arg = - begin match arg with - | None -> assert false - | Some (_, Tactics.ElimOnAnonHyp n) -> - Inv.inv_clause knd pat ids (AnonHyp n) - | Some (_, Tactics.ElimOnIdent {CAst.v=id}) -> - Inv.inv_clause knd pat ids (NamedHyp id) - | Some (_, Tactics.ElimOnConstr c) -> - let open Tactypes in - let anon = CAst.make @@ IntroNaming Namegen.IntroAnonymous in - Tactics.specialize c (Some anon) >>= fun () -> - Tacticals.New.onLastHypId (fun id -> Inv.inv_clause knd pat ids (NamedHyp id)) - end - in - on_destruction_arg inversion true (Some (None, arg)) - -let contradiction c = - let c = Option.map mk_with_bindings c in - Contradiction.contradiction c - -(** Firstorder *) - -let firstorder tac refs ids = - let open Ground_plugin in - let ids = List.map Id.to_string ids in - let tac = Option.map (fun tac -> thaw Tac2ffi.unit tac) tac in - G_ground.gen_ground_tac true tac refs ids diff --git a/src/tac2tactics.mli b/src/tac2tactics.mli deleted file mode 100644 index 026673acbf..0000000000 --- a/src/tac2tactics.mli +++ /dev/null @@ -1,124 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* intro_pattern list -> unit tactic - -val apply : advanced_flag -> evars_flag -> - constr_with_bindings thunk list -> - (Id.t * intro_pattern option) option -> unit tactic - -val induction_destruct : rec_flag -> evars_flag -> - induction_clause list -> constr_with_bindings option -> unit tactic - -val elim : evars_flag -> constr_with_bindings -> constr_with_bindings option -> - unit tactic - -val general_case_analysis : evars_flag -> constr_with_bindings -> unit tactic - -val generalize : (constr * occurrences * Name.t) list -> unit tactic - -val constructor_tac : evars_flag -> int option -> int -> bindings -> unit tactic - -val left_with_bindings : evars_flag -> bindings -> unit tactic -val right_with_bindings : evars_flag -> bindings -> unit tactic -val split_with_bindings : evars_flag -> bindings -> unit tactic - -val specialize : constr_with_bindings -> intro_pattern option -> unit tactic - -val change : Pattern.constr_pattern option -> (constr array, constr) Tac2ffi.fun1 -> clause -> unit tactic - -val rewrite : - evars_flag -> rewriting list -> clause -> unit thunk option -> unit tactic - -val symmetry : clause -> unit tactic - -val forward : bool -> unit tactic option option -> - intro_pattern option -> constr -> unit tactic - -val assert_ : assertion -> unit tactic - -val letin_pat_tac : evars_flag -> (bool * intro_pattern_naming) option -> - Name.t -> (Evd.evar_map * constr) -> clause -> unit tactic - -val reduce : Redexpr.red_expr -> clause -> unit tactic - -val simpl : GlobRef.t glob_red_flag -> - (Pattern.constr_pattern * occurrences) option -> clause -> unit tactic - -val cbv : GlobRef.t glob_red_flag -> clause -> unit tactic - -val cbn : GlobRef.t glob_red_flag -> clause -> unit tactic - -val lazy_ : GlobRef.t glob_red_flag -> clause -> unit tactic - -val unfold : (GlobRef.t * occurrences) list -> clause -> unit tactic - -val pattern : (constr * occurrences) list -> clause -> unit tactic - -val vm : (Pattern.constr_pattern * occurrences) option -> clause -> unit tactic - -val native : (Pattern.constr_pattern * occurrences) option -> clause -> unit tactic - -val eval_red : constr -> constr tactic - -val eval_hnf : constr -> constr tactic - -val eval_simpl : GlobRef.t glob_red_flag -> - (Pattern.constr_pattern * occurrences) option -> constr -> constr tactic - -val eval_cbv : GlobRef.t glob_red_flag -> constr -> constr tactic - -val eval_cbn : GlobRef.t glob_red_flag -> constr -> constr tactic - -val eval_lazy : GlobRef.t glob_red_flag -> constr -> constr tactic - -val eval_unfold : (GlobRef.t * occurrences) list -> constr -> constr tactic - -val eval_fold : constr list -> constr -> constr tactic - -val eval_pattern : (EConstr.t * occurrences) list -> constr -> constr tactic - -val eval_vm : (Pattern.constr_pattern * occurrences) option -> constr -> constr tactic - -val eval_native : (Pattern.constr_pattern * occurrences) option -> constr -> constr tactic - -val discriminate : evars_flag -> destruction_arg option -> unit tactic - -val injection : evars_flag -> intro_pattern list option -> destruction_arg option -> unit tactic - -val autorewrite : all:bool -> unit thunk option -> Id.t list -> clause -> unit tactic - -val trivial : Hints.debug -> constr thunk list -> Id.t list option -> - unit Proofview.tactic - -val auto : Hints.debug -> int option -> constr thunk list -> - Id.t list option -> unit Proofview.tactic - -val new_auto : Hints.debug -> int option -> constr thunk list -> - Id.t list option -> unit Proofview.tactic - -val eauto : Hints.debug -> int option -> int option -> constr thunk list -> - Id.t list option -> unit Proofview.tactic - -val typeclasses_eauto : Class_tactics.search_strategy option -> int option -> - Id.t list option -> unit Proofview.tactic - -val inversion : Inv.inversion_kind -> destruction_arg -> intro_pattern option -> Id.t list option -> unit tactic - -val contradiction : constr_with_bindings option -> unit tactic - -val firstorder : unit thunk option -> GlobRef.t list -> Id.t list -> unit tactic diff --git a/src/tac2types.mli b/src/tac2types.mli deleted file mode 100644 index fa31153a27..0000000000 --- a/src/tac2types.mli +++ /dev/null @@ -1,92 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* $@ - if [ $$? = 0 ]; then \ - echo " $<... OK"; \ - else \ - echo " $<... FAIL!"; \ - fi; \ - -clean: - rm -f *.log diff --git a/tests/compat.v b/tests/compat.v deleted file mode 100644 index 489fa638e4..0000000000 --- a/tests/compat.v +++ /dev/null @@ -1,58 +0,0 @@ -Require Import Ltac2.Ltac2. - -Import Ltac2.Notations. - -(** Test calls to Ltac1 from Ltac2 *) - -Ltac2 foo () := ltac1:(discriminate). - -Goal true = false -> False. -Proof. -foo (). -Qed. - -Goal true = false -> false = true. -Proof. -intros H; ltac1:(match goal with [ H : ?P |- _ ] => rewrite H end); reflexivity. -Qed. - -Goal true = false -> false = true. -Proof. -intros H; ltac1:(rewrite H); reflexivity. -Abort. - -(** Variables do not cross the compatibility layer boundary. *) -Fail Ltac2 bar nay := ltac1:(discriminate nay). - -Fail Ltac2 pose1 (v : constr) := - ltac1:(pose $v). - -(** Test calls to Ltac2 from Ltac1 *) - -Set Default Proof Mode "Classic". - -Ltac foo := ltac2:(foo ()). - -Goal true = false -> False. -Proof. -ltac2:(foo ()). -Qed. - -Goal true = false -> False. -Proof. -foo. -Qed. - -(** Variables do not cross the compatibility layer boundary. *) -Fail Ltac bar x := ltac2:(foo x). - -Ltac mytac tac := idtac "wow". - -Goal True. -Proof. -(** Fails because quotation is evaluated eagerly *) -Fail mytac ltac2:(fail). -(** One has to thunk thanks to the idtac trick *) -let t := idtac; ltac2:(fail) in mytac t. -constructor. -Qed. diff --git a/tests/errors.v b/tests/errors.v deleted file mode 100644 index c677f6af5d..0000000000 --- a/tests/errors.v +++ /dev/null @@ -1,12 +0,0 @@ -Require Import Ltac2.Ltac2. - -Goal True. -Proof. -let x := Control.plus - (fun () => let _ := constr:(nat -> 0) in 0) - (fun e => match e with Not_found => 1 | _ => 2 end) in -match Int.equal x 2 with -| true => () -| false => Control.throw (Tactic_failure None) -end. -Abort. diff --git a/tests/example1.v b/tests/example1.v deleted file mode 100644 index 023791050f..0000000000 --- a/tests/example1.v +++ /dev/null @@ -1,27 +0,0 @@ -Require Import Ltac2.Ltac2. - -Import Ltac2.Control. - -(** Alternative implementation of the hyp primitive *) -Ltac2 get_hyp_by_name x := - let h := hyps () in - let rec find x l := match l with - | [] => zero Not_found - | p :: l => - match p with - | (id, _, t) => - match Ident.equal x id with - | true => t - | false => find x l - end - end - end in - find x h. - -Print Ltac2 get_hyp_by_name. - -Goal forall n m, n + m = 0 -> n = 0. -Proof. -refine (fun () => '(fun n m H => _)). -let t := get_hyp_by_name @H in Message.print (Message.of_constr t). -Abort. diff --git a/tests/example2.v b/tests/example2.v deleted file mode 100644 index c953d25061..0000000000 --- a/tests/example2.v +++ /dev/null @@ -1,281 +0,0 @@ -Require Import Ltac2.Ltac2. - -Import Ltac2.Notations. - -Set Default Goal Selector "all". - -Goal exists n, n = 0. -Proof. -split with (x := 0). -reflexivity. -Qed. - -Goal exists n, n = 0. -Proof. -split with 0. -split. -Qed. - -Goal exists n, n = 0. -Proof. -let myvar := Std.NamedHyp @x in split with ($myvar := 0). -split. -Qed. - -Goal (forall n : nat, n = 0 -> False) -> True. -Proof. -intros H. -eelim &H. -split. -Qed. - -Goal (forall n : nat, n = 0 -> False) -> True. -Proof. -intros H. -elim &H with 0. -split. -Qed. - -Goal forall (P : nat -> Prop), (forall n m, n = m -> P n) -> P 0. -Proof. -intros P H. -Fail apply &H. -apply &H with (m := 0). -split. -Qed. - -Goal forall (P : nat -> Prop), (forall n m, n = m -> P n) -> (0 = 1) -> P 0. -Proof. -intros P H e. -apply &H with (m := 1) in e. -exact e. -Qed. - -Goal forall (P : nat -> Prop), (forall n m, n = m -> P n) -> P 0. -Proof. -intros P H. -eapply &H. -split. -Qed. - -Goal exists n, n = 0. -Proof. -Fail constructor 1. -constructor 1 with (x := 0). -split. -Qed. - -Goal exists n, n = 0. -Proof. -econstructor 1. -split. -Qed. - -Goal forall n, 0 + n = n. -Proof. -intros n. -induction &n as [|n] using nat_rect; split. -Qed. - -Goal forall n, 0 + n = n. -Proof. -intros n. -let n := @X in -let q := Std.NamedHyp @P in -induction &n as [|$n] using nat_rect with ($q := fun m => 0 + m = m); split. -Qed. - -Goal forall n, 0 + n = n. -Proof. -intros n. -destruct &n as [|n] using nat_rect; split. -Qed. - -Goal forall n, 0 + n = n. -Proof. -intros n. -let n := @X in -let q := Std.NamedHyp @P in -destruct &n as [|$n] using nat_rect with ($q := fun m => 0 + m = m); split. -Qed. - -Goal forall b1 b2, andb b1 b2 = andb b2 b1. -Proof. -intros b1 b2. -destruct &b1 as [|], &b2 as [|]; split. -Qed. - -Goal forall n m, n = 0 -> n + m = m. -Proof. -intros n m Hn. -rewrite &Hn; split. -Qed. - -Goal forall n m p, n = m -> p = m -> 0 = n -> p = 0. -Proof. -intros n m p He He' Hn. -rewrite &He, <- &He' in Hn. -rewrite &Hn. -split. -Qed. - -Goal forall n m, (m = n -> n = m) -> m = n -> n = 0 -> m = 0. -Proof. -intros n m He He' He''. -rewrite <- &He by assumption. -Control.refine (fun () => &He''). -Qed. - -Goal forall n (r := if true then n else 0), r = n. -Proof. -intros n r. -hnf in r. -split. -Qed. - -Goal 1 = 0 -> 0 = 0. -Proof. -intros H. -pattern 0 at 1. -let occ := 2 in pattern 1 at 1, 0 at $occ in H. -reflexivity. -Qed. - -Goal 1 + 1 = 2. -Proof. -vm_compute. -reflexivity. -Qed. - -Goal 1 + 1 = 2. -Proof. -native_compute. -reflexivity. -Qed. - -Goal 1 + 1 = 2 - 0 -> True. -Proof. -intros H. -vm_compute plus in H. -reflexivity. -Qed. - -Goal 1 = 0 -> True /\ True. -Proof. -intros H. -split; fold (1 + 0) (1 + 0) in H. -reflexivity. -Qed. - -Goal 1 + 1 = 2. -Proof. -cbv [ Nat.add ]. -reflexivity. -Qed. - -Goal 1 + 1 = 2. -Proof. -let x := reference:(Nat.add) in -cbn beta iota delta [ $x ]. -reflexivity. -Qed. - -Goal 1 + 1 = 2. -Proof. -simpl beta. -reflexivity. -Qed. - -Goal 1 + 1 = 2. -Proof. -lazy. -reflexivity. -Qed. - -Goal let x := 1 + 1 - 1 in x = x. -Proof. -intros x. -unfold &x at 1. -let x := reference:(Nat.sub) in unfold Nat.add, $x in x. -reflexivity. -Qed. - -Goal exists x y : nat, x = y. -Proof. -exists 0, 0; reflexivity. -Qed. - -Goal exists x y : nat, x = y. -Proof. -eexists _, 0; reflexivity. -Qed. - -Goal exists x y : nat, x = y. -Proof. -refine '(let x := 0 in _). -eexists; exists &x; reflexivity. -Qed. - -Goal True. -Proof. -pose (X := True). -constructor. -Qed. - -Goal True. -Proof. -pose True as X. -constructor. -Qed. - -Goal True. -Proof. -let x := @foo in -set ($x := True) in * |-. -constructor. -Qed. - -Goal 0 = 0. -Proof. -remember 0 as n eqn: foo at 1. -rewrite foo. -reflexivity. -Qed. - -Goal True. -Proof. -assert (H := 0 + 0). -constructor. -Qed. - -Goal True. -Proof. -assert (exists n, n = 0) as [n Hn]. -+ exists 0; reflexivity. -+ exact I. -Qed. - -Goal True -> True. -Proof. -assert (H : 0 + 0 = 0) by reflexivity. -intros x; exact x. -Qed. - -Goal 1 + 1 = 2. -Proof. -change (?a + 1 = 2) with (2 = $a + 1). -reflexivity. -Qed. - -Goal (forall n, n = 0 -> False) -> False. -Proof. -intros H. -specialize (H 0 eq_refl). -destruct H. -Qed. - -Goal (forall n, n = 0 -> False) -> False. -Proof. -intros H. -specialize (H 0 eq_refl) as []. -Qed. diff --git a/tests/matching.v b/tests/matching.v deleted file mode 100644 index 4338cbd32f..0000000000 --- a/tests/matching.v +++ /dev/null @@ -1,71 +0,0 @@ -Require Import Ltac2.Ltac2 Ltac2.Notations. - -Ltac2 Type exn ::= [ Nope ]. - -Ltac2 check_id id id' := match Ident.equal id id' with -| true => () -| false => Control.throw Nope -end. - -Goal True -> False. -Proof. -Fail -let b := { contents := true } in -let f c := - match b.(contents) with - | true => Message.print (Message.of_constr c); b.(contents) := false; fail - | false => () - end -in -(** This fails because the matching is not allowed to backtrack once - it commits to a branch*) -lazy_match! '(nat -> bool) with context [?a] => f a end. -lazy_match! Control.goal () with ?a -> ?b => Message.print (Message.of_constr b) end. - -(** This one works by taking the second match context, i.e. ?a := nat *) -let b := { contents := true } in -let f c := - match b.(contents) with - | true => b.(contents) := false; fail - | false => Message.print (Message.of_constr c) - end -in -match! '(nat -> bool) with context [?a] => f a end. -Abort. - -Goal forall (i j : unit) (x y : nat) (b : bool), True. -Proof. -Fail match! goal with -| [ h : ?t, h' : ?t |- _ ] => () -end. -intros i j x y b. -match! goal with -| [ h : ?t, h' : ?t |- _ ] => - check_id h @x; - check_id h' @y -end. -match! reverse goal with -| [ h : ?t, h' : ?t |- _ ] => - check_id h @j; - check_id h' @i -end. -Abort. - -(* Check #79 *) -Goal 2 = 3. - Control.plus - (fun () - => lazy_match! goal with - | [ |- 2 = 3 ] => Control.zero (Tactic_failure None) - | [ |- 2 = _ ] => Control.zero (Tactic_failure (Some (Message.of_string "should not be printed"))) - end) - (fun e - => match e with - | Tactic_failure c - => match c with - | None => () - | _ => Control.zero e - end - | e => Control.zero e - end). -Abort. diff --git a/tests/quot.v b/tests/quot.v deleted file mode 100644 index 624c4ad0c1..0000000000 --- a/tests/quot.v +++ /dev/null @@ -1,26 +0,0 @@ -Require Import Ltac2.Ltac2. - -(** Test for quotations *) - -Ltac2 ref0 () := reference:(&x). -Ltac2 ref1 () := reference:(nat). -Ltac2 ref2 () := reference:(Datatypes.nat). -Fail Ltac2 ref () := reference:(i_certainly_dont_exist). -Fail Ltac2 ref () := reference:(And.Me.neither). - -Goal True. -Proof. -let x := constr:(I) in -let y := constr:((fun z => z) $x) in -Control.refine (fun _ => y). -Qed. - -Goal True. -Proof. -(** Here, Ltac2 should not put its variables in the same environment as - Ltac1 otherwise the second binding fails as x is bound but not an - ident. *) -let x := constr:(I) in -let y := constr:((fun x => x) $x) in -Control.refine (fun _ => y). -Qed. diff --git a/tests/rebind.v b/tests/rebind.v deleted file mode 100644 index e1c20a2059..0000000000 --- a/tests/rebind.v +++ /dev/null @@ -1,34 +0,0 @@ -Require Import Ltac2.Ltac2 Ltac2.Notations. - -Ltac2 mutable foo () := constructor. - -Goal True. -Proof. -foo (). -Qed. - -Ltac2 Set foo := fun _ => fail. - -Goal True. -Proof. -Fail foo (). -constructor. -Qed. - -(** Not the right type *) -Fail Ltac2 Set foo := 0. - -Ltac2 bar () := (). - -(** Cannot redefine non-mutable tactics *) -Fail Ltac2 Set bar := fun _ => (). - -(** Subtype check *) - -Ltac2 mutable rec f x := f x. - -Fail Ltac2 Set f := fun x => x. - -Ltac2 mutable g x := x. - -Ltac2 Set g := f. diff --git a/tests/stuff/ltac2.v b/tests/stuff/ltac2.v deleted file mode 100644 index 370bc70d15..0000000000 --- a/tests/stuff/ltac2.v +++ /dev/null @@ -1,143 +0,0 @@ -Require Import Ltac2.Ltac2. - -Ltac2 foo (_ : int) := - let f (x : int) := x in - let _ := f 0 in - f 1. - -Print Ltac2 foo. - -Import Control. - -Ltac2 exact x := refine (fun () => x). - -Print Ltac2 refine. -Print Ltac2 exact. - -Ltac2 foo' () := ident:(bla). - -Print Ltac2 foo'. - -Ltac2 bar x h := match x with -| None => constr:(fun H => ltac2:(exact (hyp ident:(H))) -> nat) -| Some x => x -end. - -Print Ltac2 bar. - -Ltac2 qux := Some 0. - -Print Ltac2 qux. - -Ltac2 Type foo := [ Foo (int) ]. - -Fail Ltac2 qux0 := Foo None. - -Ltac2 Type 'a ref := { mutable contents : 'a }. - -Fail Ltac2 qux0 := { contents := None }. -Ltac2 foo0 () := { contents := None }. - -Print Ltac2 foo0. - -Ltac2 qux0 x := x.(contents). -Ltac2 qux1 x := x.(contents) := x.(contents). - -Ltac2 qux2 := ([1;2], true). - -Print Ltac2 qux0. -Print Ltac2 qux1. -Print Ltac2 qux2. - -Import Control. - -Ltac2 qux3 x := constr:(nat -> ltac2:(refine (fun () => hyp x))). - -Print Ltac2 qux3. - -Ltac2 Type rec nat := [ O | S (nat) ]. - -Ltac2 message_of_nat n := -let rec aux n := -match n with -| O => Message.of_string "O" -| S n => Message.concat (Message.of_string "S") (aux n) -end in aux n. - -Print Ltac2 message_of_nat. - -Ltac2 numgoals () := - let r := { contents := O } in - enter (fun () => r.(contents) := S (r.(contents))); - r.(contents). - -Print Ltac2 numgoals. - -Goal True /\ False. -Proof. -let n := numgoals () in Message.print (message_of_nat n). -refine (fun () => open_constr:((fun x => conj _ _) 0)); (). -let n := numgoals () in Message.print (message_of_nat n). - -Fail (hyp ident:(x)). -Fail (enter (fun () => hyp ident:(There_is_no_spoon); ())). - -enter (fun () => Message.print (Message.of_string "foo")). - -enter (fun () => Message.print (Message.of_constr (goal ()))). -Fail enter (fun () => Message.print (Message.of_constr (qux3 ident:(x)))). -enter (fun () => plus (fun () => constr:(_); ()) (fun _ => ())). -plus - (fun () => enter (fun () => let x := ident:(foo) in let _ := hyp x in ())) (fun _ => Message.print (Message.of_string "failed")). -let x := { contents := 0 } in -let x := x.(contents) := x.(contents) in x. -Abort. - -Ltac2 Type exn ::= [ Foo ]. - -Goal True. -Proof. -plus (fun () => zero Foo) (fun _ => ()). -Abort. - -Ltac2 Type exn ::= [ Bar (string) ]. - -Goal True. -Proof. -Fail zero (Bar "lol"). -Abort. - -Ltac2 Notation "refine!" c(thunk(constr)) := refine c. - -Goal True. -Proof. -refine! I. -Abort. - -Goal True. -Proof. -let x () := plus (fun () => 0) (fun _ => 1) in -match case x with -| Val x => - match x with - | (x, k) => Message.print (Message.of_int (k Not_found)) - end -| Err x => Message.print (Message.of_string "Err") -end. -Abort. - -Goal (forall n : nat, n = 0 -> False) -> True. -Proof. -refine (fun () => '(fun H => _)). -Std.case true (hyp @H, Std.ExplicitBindings [Std.NamedHyp @n, '0]). -refine (fun () => 'eq_refl). -Qed. - -Goal forall x, 1 + x = x + 1. -Proof. -refine (fun () => '(fun x => _)). -Std.cbv { - Std.rBeta := true; Std.rMatch := true; Std.rFix := true; Std.rCofix := true; - Std.rZeta := true; Std.rDelta := true; Std.rConst := []; -} { Std.on_hyps := None; Std.on_concl := Std.AllOccurrences }. -Abort. diff --git a/tests/tacticals.v b/tests/tacticals.v deleted file mode 100644 index 1a2fbcbb37..0000000000 --- a/tests/tacticals.v +++ /dev/null @@ -1,34 +0,0 @@ -Require Import Ltac2.Ltac2. - -Import Ltac2.Notations. - -Goal True. -Proof. -Fail fail. -Fail solve [ () ]. -try fail. -repeat fail. -repeat (). -solve [ constructor ]. -Qed. - -Goal True. -Proof. -first [ - Message.print (Message.of_string "Yay"); fail -| constructor -| Message.print (Message.of_string "I won't be printed") -]. -Qed. - -Goal True /\ True. -Proof. -Fail split > [ split | |]. -split > [split | split]. -Qed. - -Goal True /\ (True -> True) /\ True. -Proof. -split > [ | split] > [split | .. | split]. -intros H; refine &H. -Qed. diff --git a/tests/typing.v b/tests/typing.v deleted file mode 100644 index 9f18292716..0000000000 --- a/tests/typing.v +++ /dev/null @@ -1,72 +0,0 @@ -Require Import Ltac2.Ltac2. - -(** Ltac2 is typed à la ML. *) - -Ltac2 test0 n := Int.add n 1. - -Print Ltac2 test0. - -Ltac2 test1 () := test0 0. - -Print Ltac2 test1. - -Fail Ltac2 test2 () := test0 true. - -Fail Ltac2 test2 () := test0 0 0. - -Ltac2 test3 f x := x, (f x, x). - -Print Ltac2 test3. - -(** Polymorphism *) - -Ltac2 rec list_length l := -match l with -| [] => 0 -| x :: l => Int.add 1 (list_length l) -end. - -Print Ltac2 list_length. - -(** Pattern-matching *) - -Ltac2 ifb b f g := match b with -| true => f () -| false => g () -end. - -Print Ltac2 ifb. - -Ltac2 if_not_found e f g := match e with -| Not_found => f () -| _ => g () -end. - -Fail Ltac2 ifb' b f g := match b with -| true => f () -end. - -Fail Ltac2 if_not_found' e f g := match e with -| Not_found => f () -end. - -(** Reimplementing 'do'. Return value of the function useless. *) - -Ltac2 rec do n tac := match Int.equal n 0 with -| true => () -| false => tac (); do (Int.sub n 1) tac -end. - -Print Ltac2 do. - -(** Non-function pure values are OK. *) - -Ltac2 tuple0 := ([1; 2], true, (fun () => "yay")). - -Print Ltac2 tuple0. - -(** Impure values are not. *) - -Fail Ltac2 not_a_value := { contents := 0 }. -Fail Ltac2 not_a_value := "nope". -Fail Ltac2 not_a_value := list_length []. diff --git a/theories/Array.v b/theories/Array.v deleted file mode 100644 index 11b64e3515..0000000000 --- a/theories/Array.v +++ /dev/null @@ -1,14 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* 'a -> 'a array := "ltac2" "array_make". -Ltac2 @external length : 'a array -> int := "ltac2" "array_length". -Ltac2 @external get : 'a array -> int -> 'a := "ltac2" "array_get". -Ltac2 @external set : 'a array -> int -> 'a -> unit := "ltac2" "array_set". diff --git a/theories/Char.v b/theories/Char.v deleted file mode 100644 index 29fef60f2c..0000000000 --- a/theories/Char.v +++ /dev/null @@ -1,12 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* char := "ltac2" "char_of_int". -Ltac2 @external to_int : char -> int := "ltac2" "char_to_int". diff --git a/theories/Constr.v b/theories/Constr.v deleted file mode 100644 index d8d222730e..0000000000 --- a/theories/Constr.v +++ /dev/null @@ -1,72 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* constr := "ltac2" "constr_type". -(** Return the type of a term *) - -Ltac2 @ external equal : constr -> constr -> bool := "ltac2" "constr_equal". -(** Strict syntactic equality: only up to α-conversion and evar expansion *) - -Module Unsafe. - -(** Low-level access to kernel terms. Use with care! *) - -Ltac2 Type case. - -Ltac2 Type kind := [ -| Rel (int) -| Var (ident) -| Meta (meta) -| Evar (evar, constr array) -| Sort (sort) -| Cast (constr, cast, constr) -| Prod (ident option, constr, constr) -| Lambda (ident option, constr, constr) -| LetIn (ident option, constr, constr, constr) -| App (constr, constr array) -| Constant (constant, instance) -| Ind (inductive, instance) -| Constructor (constructor, instance) -| Case (case, constr, constr, constr array) -| Fix (int array, int, ident option array, constr array, constr array) -| CoFix (int, ident option array, constr array, constr array) -| Proj (projection, constr) -]. - -Ltac2 @ external kind : constr -> kind := "ltac2" "constr_kind". - -Ltac2 @ external make : kind -> constr := "ltac2" "constr_make". - -Ltac2 @ external check : constr -> constr result := "ltac2" "constr_check". -(** Checks that a constr generated by unsafe means is indeed safe in the - current environment, and returns it, or the error otherwise. Panics if - not focussed. *) - -Ltac2 @ external substnl : constr list -> int -> constr -> constr := "ltac2" "constr_substnl". -(** [substnl [r₁;...;rₙ] k c] substitutes in parallel [Rel(k+1); ...; Rel(k+n)] with - [r₁;...;rₙ] in [c]. *) - -Ltac2 @ external closenl : ident list -> int -> constr -> constr := "ltac2" "constr_closenl". -(** [closenl [x₁;...;xₙ] k c] abstracts over variables [x₁;...;xₙ] and replaces them with - [Rel(k); ...; Rel(k+n-1)] in [c]. If two names are identical, the one of least index is kept. *) - -Ltac2 @ external case : inductive -> case := "ltac2" "constr_case". -(** Generate the case information for a given inductive type. *) - -Ltac2 @ external constructor : inductive -> int -> constructor := "ltac2" "constr_constructor". -(** Generate the i-th constructor for a given inductive type. Indexing starts - at 0. Panics if there is no such constructor. *) - -End Unsafe. - -Ltac2 @ external in_context : ident -> constr -> (unit -> unit) -> constr := "ltac2" "constr_in_context". -(** On a focussed goal [Γ ⊢ A], [in_context id c tac] evaluates [tac] in a - focussed goal [Γ, id : c ⊢ ?X] and returns [fun (id : c) => t] where [t] is - the proof built by the tactic. *) diff --git a/theories/Control.v b/theories/Control.v deleted file mode 100644 index 071c2ea8ce..0000000000 --- a/theories/Control.v +++ /dev/null @@ -1,76 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* 'a := "ltac2" "throw". -(** Fatal exception throwing. This does not induce backtracking. *) - -(** Generic backtracking control *) - -Ltac2 @ external zero : exn -> 'a := "ltac2" "zero". -Ltac2 @ external plus : (unit -> 'a) -> (exn -> 'a) -> 'a := "ltac2" "plus". -Ltac2 @ external once : (unit -> 'a) -> 'a := "ltac2" "once". -Ltac2 @ external dispatch : (unit -> unit) list -> unit := "ltac2" "dispatch". -Ltac2 @ external extend : (unit -> unit) list -> (unit -> unit) -> (unit -> unit) list -> unit := "ltac2" "extend". -Ltac2 @ external enter : (unit -> unit) -> unit := "ltac2" "enter". -Ltac2 @ external case : (unit -> 'a) -> ('a * (exn -> 'a)) result := "ltac2" "case". - -(** Proof state manipulation *) - -Ltac2 @ external focus : int -> int -> (unit -> 'a) -> 'a := "ltac2" "focus". -Ltac2 @ external shelve : unit -> unit := "ltac2" "shelve". -Ltac2 @ external shelve_unifiable : unit -> unit := "ltac2" "shelve_unifiable". - -Ltac2 @ external new_goal : evar -> unit := "ltac2" "new_goal". -(** Adds the given evar to the list of goals as the last one. If it is - already defined in the current state, don't do anything. Panics if the - evar is not in the current state. *) - -Ltac2 @ external progress : (unit -> 'a) -> 'a := "ltac2" "progress". - -(** Goal inspection *) - -Ltac2 @ external goal : unit -> constr := "ltac2" "goal". -(** Panics if there is not exactly one goal under focus. Otherwise returns - the conclusion of this goal. *) - -Ltac2 @ external hyp : ident -> constr := "ltac2" "hyp". -(** Panics if there is more than one goal under focus. If there is no - goal under focus, looks for the section variable with the given name. - If there is one, looks for the hypothesis with the given name. *) - -Ltac2 @ external hyps : unit -> (ident * constr option * constr) list := "ltac2" "hyps". -(** Panics if there is more than one goal under focus. If there is no - goal under focus, returns the list of section variables. - If there is one, returns the list of hypotheses. In both cases, the - list is ordered with rightmost values being last introduced. *) - -(** Refinement *) - -Ltac2 @ external refine : (unit -> constr) -> unit := "ltac2" "refine". - -(** Evars *) - -Ltac2 @ external with_holes : (unit -> 'a) -> ('a -> 'b) -> 'b := "ltac2" "with_holes". -(** [with_holes x f] evaluates [x], then apply [f] to the result, and fails if - all evars generated by the call to [x] have not been solved when [f] - returns. *) - -(** Misc *) - -Ltac2 @ external time : string option -> (unit -> 'a) -> 'a := "ltac2" "time". -(** Displays the time taken by a tactic to evaluate. *) - -Ltac2 @ external abstract : ident option -> (unit -> unit) -> unit := "ltac2" "abstract". -(** Abstract a subgoal. *) - -Ltac2 @ external check_interrupt : unit -> unit := "ltac2" "check_interrupt". -(** For internal use. *) diff --git a/theories/Env.v b/theories/Env.v deleted file mode 100644 index c9b250f4ba..0000000000 --- a/theories/Env.v +++ /dev/null @@ -1,27 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* Std.reference option := "ltac2" "env_get". -(** Returns the global reference corresponding to the absolute name given as - argument if it exists. *) - -Ltac2 @ external expand : ident list -> Std.reference list := "ltac2" "env_expand". -(** Returns the list of all global references whose absolute name contains - the argument list as a prefix. *) - -Ltac2 @ external path : Std.reference -> ident list := "ltac2" "env_path". -(** Returns the absolute name of the given reference. Panics if the reference - does not exist. *) - -Ltac2 @ external instantiate : Std.reference -> constr := "ltac2" "env_instantiate". -(** Returns a fresh instance of the corresponding reference, in particular - generating fresh universe variables and constraints when this reference is - universe-polymorphic. *) diff --git a/theories/Fresh.v b/theories/Fresh.v deleted file mode 100644 index 5e876bb077..0000000000 --- a/theories/Fresh.v +++ /dev/null @@ -1,26 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* t -> t := "ltac2" "fresh_free_union". - -Ltac2 @ external of_ids : ident list -> t := "ltac2" "fresh_free_of_ids". - -Ltac2 @ external of_constr : constr -> t := "ltac2" "fresh_free_of_constr". - -End Free. - -Ltac2 @ external fresh : Free.t -> ident -> ident := "ltac2" "fresh_fresh". -(** Generate a fresh identifier with the given base name which is not a - member of the provided set of free variables. *) diff --git a/theories/Ident.v b/theories/Ident.v deleted file mode 100644 index 55456afbe2..0000000000 --- a/theories/Ident.v +++ /dev/null @@ -1,17 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* t -> bool := "ltac2" "ident_equal". - -Ltac2 @ external of_string : string -> t option := "ltac2" "ident_of_string". - -Ltac2 @ external to_string : t -> string := "ltac2" "ident_to_string". diff --git a/theories/Init.v b/theories/Init.v deleted file mode 100644 index 16e7d7a6f9..0000000000 --- a/theories/Init.v +++ /dev/null @@ -1,69 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* int -> bool := "ltac2" "int_equal". -Ltac2 @ external compare : int -> int -> int := "ltac2" "int_compare". -Ltac2 @ external add : int -> int -> int := "ltac2" "int_add". -Ltac2 @ external sub : int -> int -> int := "ltac2" "int_sub". -Ltac2 @ external mul : int -> int -> int := "ltac2" "int_mul". -Ltac2 @ external neg : int -> int := "ltac2" "int_neg". diff --git a/theories/Ltac1.v b/theories/Ltac1.v deleted file mode 100644 index c4e0b606d0..0000000000 --- a/theories/Ltac1.v +++ /dev/null @@ -1,36 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* t := "ltac2" "ltac1_ref". -(** Returns the Ltac1 definition with the given absolute name. *) - -Ltac2 @ external run : t -> unit := "ltac2" "ltac1_run". -(** Runs an Ltac1 value, assuming it is a 'tactic', i.e. not returning - anything. *) - -Ltac2 @ external apply : t -> t list -> (t -> unit) -> unit := "ltac2" "ltac1_apply". -(** Applies an Ltac1 value to a list of arguments, and provides the result in - CPS style. It does **not** run the returned value. *) - -(** Conversion functions *) - -Ltac2 @ external of_constr : constr -> t := "ltac2" "ltac1_of_constr". -Ltac2 @ external to_constr : t -> constr option := "ltac2" "ltac1_to_constr". - -Ltac2 @ external of_list : t list -> t := "ltac2" "ltac1_of_list". -Ltac2 @ external to_list : t -> t list option := "ltac2" "ltac1_to_list". diff --git a/theories/Ltac2.v b/theories/Ltac2.v deleted file mode 100644 index ac90f63560..0000000000 --- a/theories/Ltac2.v +++ /dev/null @@ -1,24 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* unit := "ltac2" "print". - -Ltac2 @ external of_string : string -> message := "ltac2" "message_of_string". - -Ltac2 @ external of_int : int -> message := "ltac2" "message_of_int". - -Ltac2 @ external of_ident : ident -> message := "ltac2" "message_of_ident". - -Ltac2 @ external of_constr : constr -> message := "ltac2" "message_of_constr". -(** Panics if there is more than one goal under focus. *) - -Ltac2 @ external of_exn : exn -> message := "ltac2" "message_of_exn". -(** Panics if there is more than one goal under focus. *) - -Ltac2 @ external concat : message -> message -> message := "ltac2" "message_concat". diff --git a/theories/Notations.v b/theories/Notations.v deleted file mode 100644 index f4621656d6..0000000000 --- a/theories/Notations.v +++ /dev/null @@ -1,568 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* f e -| Val ans => - let (x, k) := ans in - Control.plus (fun _ => x) k -end. - -Ltac2 ifcatch t s f := -match Control.case t with -| Err e => f e -| Val ans => - let (x, k) := ans in - Control.plus (fun _ => s x) (fun e => s (k e)) -end. - -Ltac2 fail0 (_ : unit) := Control.enter (fun _ => Control.zero (Tactic_failure None)). - -Ltac2 Notation fail := fail0 (). - -Ltac2 try0 t := Control.enter (fun _ => orelse t (fun _ => ())). - -Ltac2 Notation try := try0. - -Ltac2 rec repeat0 (t : unit -> unit) := - Control.enter (fun () => - ifcatch (fun _ => Control.progress t) - (fun _ => Control.check_interrupt (); repeat0 t) (fun _ => ())). - -Ltac2 Notation repeat := repeat0. - -Ltac2 dispatch0 t (head, tail) := - match tail with - | None => Control.enter (fun _ => t (); Control.dispatch head) - | Some tacs => - let (def, rem) := tacs in - Control.enter (fun _ => t (); Control.extend head def rem) - end. - -Ltac2 Notation t(thunk(self)) ">" "[" l(dispatch) "]" : 4 := dispatch0 t l. - -Ltac2 do0 n t := - let rec aux n t := match Int.equal n 0 with - | true => () - | false => t (); aux (Int.sub n 1) t - end in - aux (n ()) t. - -Ltac2 Notation do := do0. - -Ltac2 Notation once := Control.once. - -Ltac2 progress0 tac := Control.enter (fun _ => Control.progress tac). - -Ltac2 Notation progress := progress0. - -Ltac2 rec first0 tacs := -match tacs with -| [] => Control.zero (Tactic_failure None) -| tac :: tacs => Control.enter (fun _ => orelse tac (fun _ => first0 tacs)) -end. - -Ltac2 Notation "first" "[" tacs(list0(thunk(tactic(6)), "|")) "]" := first0 tacs. - -Ltac2 complete tac := - let ans := tac () in - Control.enter (fun () => Control.zero (Tactic_failure None)); - ans. - -Ltac2 rec solve0 tacs := -match tacs with -| [] => Control.zero (Tactic_failure None) -| tac :: tacs => - Control.enter (fun _ => orelse (fun _ => complete tac) (fun _ => solve0 tacs)) -end. - -Ltac2 Notation "solve" "[" tacs(list0(thunk(tactic(6)), "|")) "]" := solve0 tacs. - -Ltac2 time0 tac := Control.time None tac. - -Ltac2 Notation time := time0. - -Ltac2 abstract0 tac := Control.abstract None tac. - -Ltac2 Notation abstract := abstract0. - -(** Base tactics *) - -(** Note that we redeclare notations that can be parsed as mere identifiers - as abbreviations, so that it allows to parse them as function arguments - without having to write them within parentheses. *) - -(** Enter and check evar resolution *) -Ltac2 enter_h ev f arg := -match ev with -| true => Control.enter (fun () => f ev (arg ())) -| false => - Control.enter (fun () => - Control.with_holes arg (fun x => f ev x)) -end. - -Ltac2 intros0 ev p := - Control.enter (fun () => Std.intros false p). - -Ltac2 Notation "intros" p(intropatterns) := intros0 false p. -Ltac2 Notation intros := intros. - -Ltac2 Notation "eintros" p(intropatterns) := intros0 true p. -Ltac2 Notation eintros := eintros. - -Ltac2 split0 ev bnd := - enter_h ev Std.split bnd. - -Ltac2 Notation "split" bnd(thunk(with_bindings)) := split0 false bnd. -Ltac2 Notation split := split. - -Ltac2 Notation "esplit" bnd(thunk(with_bindings)) := split0 true bnd. -Ltac2 Notation esplit := esplit. - -Ltac2 exists0 ev bnds := match bnds with -| [] => split0 ev (fun () => Std.NoBindings) -| _ => - let rec aux bnds := match bnds with - | [] => () - | bnd :: bnds => split0 ev bnd; aux bnds - end in - aux bnds -end. - -Ltac2 Notation "exists" bnd(list0(thunk(bindings), ",")) := exists0 false bnd. -(* Ltac2 Notation exists := exists. *) - -Ltac2 Notation "eexists" bnd(list0(thunk(bindings), ",")) := exists0 true bnd. -Ltac2 Notation eexists := eexists. - -Ltac2 left0 ev bnd := enter_h ev Std.left bnd. - -Ltac2 Notation "left" bnd(thunk(with_bindings)) := left0 false bnd. -Ltac2 Notation left := left. - -Ltac2 Notation "eleft" bnd(thunk(with_bindings)) := left0 true bnd. -Ltac2 Notation eleft := eleft. - -Ltac2 right0 ev bnd := enter_h ev Std.right bnd. - -Ltac2 Notation "right" bnd(thunk(with_bindings)) := right0 false bnd. -Ltac2 Notation right := right. - -Ltac2 Notation "eright" bnd(thunk(with_bindings)) := right0 true bnd. -Ltac2 Notation eright := eright. - -Ltac2 constructor0 ev n bnd := - enter_h ev (fun ev bnd => Std.constructor_n ev n bnd) bnd. - -Ltac2 Notation "constructor" := Control.enter (fun () => Std.constructor false). -Ltac2 Notation constructor := constructor. -Ltac2 Notation "constructor" n(tactic) bnd(thunk(with_bindings)) := constructor0 false n bnd. - -Ltac2 Notation "econstructor" := Control.enter (fun () => Std.constructor true). -Ltac2 Notation econstructor := econstructor. -Ltac2 Notation "econstructor" n(tactic) bnd(thunk(with_bindings)) := constructor0 true n bnd. - -Ltac2 specialize0 c pat := - enter_h false (fun _ c => Std.specialize c pat) c. - -Ltac2 Notation "specialize" c(thunk(seq(constr, with_bindings))) ipat(opt(seq("as", intropattern))) := - specialize0 c ipat. - -Ltac2 elim0 ev c bnd use := - let f ev (c, bnd, use) := Std.elim ev (c, bnd) use in - enter_h ev f (fun () => c (), bnd (), use ()). - -Ltac2 Notation "elim" c(thunk(constr)) bnd(thunk(with_bindings)) - use(thunk(opt(seq("using", constr, with_bindings)))) := - elim0 false c bnd use. - -Ltac2 Notation "eelim" c(thunk(constr)) bnd(thunk(with_bindings)) - use(thunk(opt(seq("using", constr, with_bindings)))) := - elim0 true c bnd use. - -Ltac2 apply0 adv ev cb cl := - Std.apply adv ev cb cl. - -Ltac2 Notation "eapply" - cb(list1(thunk(seq(constr, with_bindings)), ",")) - cl(opt(seq("in", ident, opt(seq("as", intropattern))))) := - apply0 true true cb cl. - -Ltac2 Notation "apply" - cb(list1(thunk(seq(constr, with_bindings)), ",")) - cl(opt(seq("in", ident, opt(seq("as", intropattern))))) := - apply0 true false cb cl. - -Ltac2 default_on_concl cl := -match cl with -| None => { Std.on_hyps := Some []; Std.on_concl := Std.AllOccurrences } -| Some cl => cl -end. - -Ltac2 pose0 ev p := - enter_h ev (fun ev (na, p) => Std.pose na p) p. - -Ltac2 Notation "pose" p(thunk(pose)) := - pose0 false p. - -Ltac2 Notation "epose" p(thunk(pose)) := - pose0 true p. - -Ltac2 Notation "set" p(thunk(pose)) cl(opt(clause)) := - Std.set false p (default_on_concl cl). - -Ltac2 Notation "eset" p(thunk(pose)) cl(opt(clause)) := - Std.set true p (default_on_concl cl). - -Ltac2 assert0 ev ast := - enter_h ev (fun _ ast => Std.assert ast) ast. - -Ltac2 Notation "assert" ast(thunk(assert)) := assert0 false ast. - -Ltac2 Notation "eassert" ast(thunk(assert)) := assert0 true ast. - -Ltac2 default_everywhere cl := -match cl with -| None => { Std.on_hyps := None; Std.on_concl := Std.AllOccurrences } -| Some cl => cl -end. - -Ltac2 Notation "remember" - c(thunk(open_constr)) - na(opt(seq("as", ident))) - pat(opt(seq("eqn", ":", intropattern))) - cl(opt(clause)) := - Std.remember false na c pat (default_everywhere cl). - -Ltac2 Notation "eremember" - c(thunk(open_constr)) - na(opt(seq("as", ident))) - pat(opt(seq("eqn", ":", intropattern))) - cl(opt(clause)) := - Std.remember true na c pat (default_everywhere cl). - -Ltac2 induction0 ev ic use := - let f ev use := Std.induction ev ic use in - enter_h ev f use. - -Ltac2 Notation "induction" - ic(list1(induction_clause, ",")) - use(thunk(opt(seq("using", constr, with_bindings)))) := - induction0 false ic use. - -Ltac2 Notation "einduction" - ic(list1(induction_clause, ",")) - use(thunk(opt(seq("using", constr, with_bindings)))) := - induction0 true ic use. - -Ltac2 generalize0 gen := - enter_h false (fun _ gen => Std.generalize gen) gen. - -Ltac2 Notation "generalize" - gen(thunk(list1(seq (open_constr, occurrences, opt(seq("as", ident))), ","))) := - generalize0 gen. - -Ltac2 destruct0 ev ic use := - let f ev use := Std.destruct ev ic use in - enter_h ev f use. - -Ltac2 Notation "destruct" - ic(list1(induction_clause, ",")) - use(thunk(opt(seq("using", constr, with_bindings)))) := - destruct0 false ic use. - -Ltac2 Notation "edestruct" - ic(list1(induction_clause, ",")) - use(thunk(opt(seq("using", constr, with_bindings)))) := - destruct0 true ic use. - -Ltac2 Notation "simple" "inversion" - arg(destruction_arg) - pat(opt(seq("as", intropattern))) - ids(opt(seq("in", list1(ident)))) := - Std.inversion Std.SimpleInversion arg pat ids. - -Ltac2 Notation "inversion" - arg(destruction_arg) - pat(opt(seq("as", intropattern))) - ids(opt(seq("in", list1(ident)))) := - Std.inversion Std.FullInversion arg pat ids. - -Ltac2 Notation "inversion_clear" - arg(destruction_arg) - pat(opt(seq("as", intropattern))) - ids(opt(seq("in", list1(ident)))) := - Std.inversion Std.FullInversionClear arg pat ids. - -Ltac2 Notation "red" cl(opt(clause)) := - Std.red (default_on_concl cl). -Ltac2 Notation red := red. - -Ltac2 Notation "hnf" cl(opt(clause)) := - Std.hnf (default_on_concl cl). -Ltac2 Notation hnf := hnf. - -Ltac2 Notation "simpl" s(strategy) pl(opt(seq(pattern, occurrences))) cl(opt(clause)) := - Std.simpl s pl (default_on_concl cl). -Ltac2 Notation simpl := simpl. - -Ltac2 Notation "cbv" s(strategy) cl(opt(clause)) := - Std.cbv s (default_on_concl cl). -Ltac2 Notation cbv := cbv. - -Ltac2 Notation "cbn" s(strategy) cl(opt(clause)) := - Std.cbn s (default_on_concl cl). -Ltac2 Notation cbn := cbn. - -Ltac2 Notation "lazy" s(strategy) cl(opt(clause)) := - Std.lazy s (default_on_concl cl). -Ltac2 Notation lazy := lazy. - -Ltac2 Notation "unfold" pl(list1(seq(reference, occurrences), ",")) cl(opt(clause)) := - Std.unfold pl (default_on_concl cl). - -Ltac2 fold0 pl cl := - let cl := default_on_concl cl in - Control.enter (fun () => Control.with_holes pl (fun pl => Std.fold pl cl)). - -Ltac2 Notation "fold" pl(thunk(list1(open_constr))) cl(opt(clause)) := - fold0 pl cl. - -Ltac2 Notation "pattern" pl(list1(seq(constr, occurrences), ",")) cl(opt(clause)) := - Std.pattern pl (default_on_concl cl). - -Ltac2 Notation "vm_compute" pl(opt(seq(pattern, occurrences))) cl(opt(clause)) := - Std.vm pl (default_on_concl cl). -Ltac2 Notation vm_compute := vm_compute. - -Ltac2 Notation "native_compute" pl(opt(seq(pattern, occurrences))) cl(opt(clause)) := - Std.native pl (default_on_concl cl). -Ltac2 Notation native_compute := native_compute. - -Ltac2 change0 p cl := - let (pat, c) := p in - Std.change pat c (default_on_concl cl). - -Ltac2 Notation "change" c(conversion) cl(opt(clause)) := change0 c cl. - -Ltac2 rewrite0 ev rw cl tac := - let cl := default_on_concl cl in - Std.rewrite ev rw cl tac. - -Ltac2 Notation "rewrite" - rw(list1(rewriting, ",")) - cl(opt(clause)) - tac(opt(seq("by", thunk(tactic)))) := - rewrite0 false rw cl tac. - -Ltac2 Notation "erewrite" - rw(list1(rewriting, ",")) - cl(opt(clause)) - tac(opt(seq("by", thunk(tactic)))) := - rewrite0 true rw cl tac. - -(** coretactics *) - -Ltac2 exact0 ev c := - Control.enter (fun _ => - match ev with - | true => - let c := c () in - Control.refine (fun _ => c) - | false => - Control.with_holes c (fun c => Control.refine (fun _ => c)) - end - ). - -Ltac2 Notation "exact" c(thunk(open_constr)) := exact0 false c. -Ltac2 Notation "eexact" c(thunk(open_constr)) := exact0 true c. - -Ltac2 Notation "intro" id(opt(ident)) mv(opt(move_location)) := Std.intro id mv. -Ltac2 Notation intro := intro. - -Ltac2 Notation "move" id(ident) mv(move_location) := Std.move id mv. - -Ltac2 Notation reflexivity := Std.reflexivity (). - -Ltac2 symmetry0 cl := - Std.symmetry (default_on_concl cl). - -Ltac2 Notation "symmetry" cl(opt(clause)) := symmetry0 cl. -Ltac2 Notation symmetry := symmetry. - -Ltac2 Notation "revert" ids(list1(ident)) := Std.revert ids. - -Ltac2 Notation assumption := Std.assumption (). - -Ltac2 Notation etransitivity := Std.etransitivity (). - -Ltac2 Notation admit := Std.admit (). - -Ltac2 clear0 ids := match ids with -| [] => Std.keep [] -| _ => Std.clear ids -end. - -Ltac2 Notation "clear" ids(list0(ident)) := clear0 ids. -Ltac2 Notation "clear" "-" ids(list1(ident)) := Std.keep ids. -Ltac2 Notation clear := clear. - -Ltac2 Notation refine := Control.refine. - -(** extratactics *) - -Ltac2 absurd0 c := Control.enter (fun _ => Std.absurd (c ())). - -Ltac2 Notation "absurd" c(thunk(open_constr)) := absurd0 c. - -Ltac2 subst0 ids := match ids with -| [] => Std.subst_all () -| _ => Std.subst ids -end. - -Ltac2 Notation "subst" ids(list0(ident)) := subst0 ids. -Ltac2 Notation subst := subst. - -Ltac2 Notation "discriminate" arg(opt(destruction_arg)) := - Std.discriminate false arg. -Ltac2 Notation discriminate := discriminate. - -Ltac2 Notation "ediscriminate" arg(opt(destruction_arg)) := - Std.discriminate true arg. -Ltac2 Notation ediscriminate := ediscriminate. - -Ltac2 Notation "injection" arg(opt(destruction_arg)) ipat(opt(seq("as", intropatterns))):= - Std.injection false ipat arg. - -Ltac2 Notation "einjection" arg(opt(destruction_arg)) ipat(opt(seq("as", intropatterns))):= - Std.injection true ipat arg. - -(** Auto *) - -Ltac2 default_db dbs := match dbs with -| None => Some [] -| Some dbs => - match dbs with - | None => None - | Some l => Some l - end -end. - -Ltac2 default_list use := match use with -| None => [] -| Some use => use -end. - -Ltac2 trivial0 use dbs := - let dbs := default_db dbs in - let use := default_list use in - Std.trivial Std.Off use dbs. - -Ltac2 Notation "trivial" - use(opt(seq("using", list1(thunk(constr), ",")))) - dbs(opt(seq("with", hintdb))) := trivial0 use dbs. - -Ltac2 Notation trivial := trivial. - -Ltac2 auto0 n use dbs := - let dbs := default_db dbs in - let use := default_list use in - Std.auto Std.Off n use dbs. - -Ltac2 Notation "auto" n(opt(tactic(0))) - use(opt(seq("using", list1(thunk(constr), ",")))) - dbs(opt(seq("with", hintdb))) := auto0 n use dbs. - -Ltac2 Notation auto := auto. - -Ltac2 new_eauto0 n use dbs := - let dbs := default_db dbs in - let use := default_list use in - Std.new_auto Std.Off n use dbs. - -Ltac2 Notation "new" "auto" n(opt(tactic(0))) - use(opt(seq("using", list1(thunk(constr), ",")))) - dbs(opt(seq("with", hintdb))) := new_eauto0 n use dbs. - -Ltac2 eauto0 n p use dbs := - let dbs := default_db dbs in - let use := default_list use in - Std.eauto Std.Off n p use dbs. - -Ltac2 Notation "eauto" n(opt(tactic(0))) p(opt(tactic(0))) - use(opt(seq("using", list1(thunk(constr), ",")))) - dbs(opt(seq("with", hintdb))) := eauto0 n p use dbs. - -Ltac2 Notation eauto := eauto. - -Ltac2 Notation "typeclasses_eauto" n(opt(tactic(0))) - dbs(opt(seq("with", list1(ident)))) := Std.typeclasses_eauto None n dbs. - -Ltac2 Notation "typeclasses_eauto" "bfs" n(opt(tactic(0))) - dbs(opt(seq("with", list1(ident)))) := Std.typeclasses_eauto (Some Std.BFS) n dbs. - -Ltac2 Notation typeclasses_eauto := typeclasses_eauto. - -(** Congruence *) - -Ltac2 f_equal0 () := ltac1:(f_equal). -Ltac2 Notation f_equal := f_equal0 (). - -(** Firstorder *) - -Ltac2 firstorder0 tac refs ids := - let refs := default_list refs in - let ids := default_list ids in - Std.firstorder tac refs ids. - -Ltac2 Notation "firstorder" - tac(opt(thunk(tactic))) - refs(opt(seq("using", list1(reference, ",")))) - ids(opt(seq("with", list1(ident)))) := firstorder0 tac refs ids. - -(** now *) - -Ltac2 now0 t := t (); ltac1:(easy). -Ltac2 Notation "now" t(thunk(self)) := now0 t. diff --git a/theories/Pattern.v b/theories/Pattern.v deleted file mode 100644 index 8d1fb0cd8a..0000000000 --- a/theories/Pattern.v +++ /dev/null @@ -1,145 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* context := - "ltac2" "pattern_empty_context". -(** A trivial context only made of the hole. *) - -Ltac2 @ external matches : t -> constr -> (ident * constr) list := - "ltac2" "pattern_matches". -(** If the term matches the pattern, returns the bound variables. If it doesn't, - fail with [Match_failure]. Panics if not focussed. *) - -Ltac2 @ external matches_subterm : t -> constr -> context * ((ident * constr) list) := - "ltac2" "pattern_matches_subterm". -(** Returns a stream of results corresponding to all of the subterms of the term - that matches the pattern as in [matches]. The stream is encoded as a - backtracking value whose last exception is [Match_failure]. The additional - value compared to [matches] is the context of the match, to be filled with - the instantiate function. *) - -Ltac2 @ external matches_vect : t -> constr -> constr array := - "ltac2" "pattern_matches_vect". -(** Internal version of [matches] that does not return the identifiers. *) - -Ltac2 @ external matches_subterm_vect : t -> constr -> context * constr array := - "ltac2" "pattern_matches_subterm_vect". -(** Internal version of [matches_subterms] that does not return the identifiers. *) - -Ltac2 @ external matches_goal : bool -> (match_kind * t) list -> (match_kind * t) -> - ident array * context array * constr array * context := - "ltac2" "pattern_matches_goal". -(** Given a list of patterns [hpats] for hypotheses and one pattern [cpat] for the - conclusion, [matches_goal rev hpats cpat] produces (a stream of) tuples of: - - An array of idents, whose size is the length of [hpats], corresponding to the - name of matched hypotheses. - - An array of contexts, whose size is the length of [hpats], corresponding to - the contexts matched for every hypothesis pattern. In case the match kind of - a hypothesis was [MatchPattern], the corresponding context is ensured to be empty. - - An array of terms, whose size is the total number of pattern variables without - duplicates. Terms are ordered by identifier order, e.g. ?a comes before ?b. - - A context corresponding to the conclusion, which is ensured to be empty if - the kind of [cpat] was [MatchPattern]. - This produces a backtracking stream of results containing all the possible - result combinations. The order of considered hypotheses is reversed if [rev] - is true. -*) - -Ltac2 @ external instantiate : context -> constr -> constr := - "ltac2" "pattern_instantiate". -(** Fill the hole of a context with the given term. *) - -(** Implementation of Ltac matching over terms and goals *) - -Ltac2 lazy_match0 t pats := - let rec interp m := match m with - | [] => Control.zero Match_failure - | p :: m => - let next _ := interp m in - let (knd, pat, f) := p in - let p := match knd with - | MatchPattern => - (fun _ => - let context := empty_context () in - let bind := matches_vect pat t in - fun _ => f context bind) - | MatchContext => - (fun _ => - let (context, bind) := matches_subterm_vect pat t in - fun _ => f context bind) - end in - Control.plus p next - end in - Control.once (fun () => interp pats) (). - -Ltac2 multi_match0 t pats := - let rec interp m := match m with - | [] => Control.zero Match_failure - | p :: m => - let next _ := interp m in - let (knd, pat, f) := p in - let p := match knd with - | MatchPattern => - (fun _ => - let context := empty_context () in - let bind := matches_vect pat t in - f context bind) - | MatchContext => - (fun _ => - let (context, bind) := matches_subterm_vect pat t in - f context bind) - end in - Control.plus p next - end in - interp pats. - -Ltac2 one_match0 t m := Control.once (fun _ => multi_match0 t m). - -Ltac2 lazy_goal_match0 rev pats := - let rec interp m := match m with - | [] => Control.zero Match_failure - | p :: m => - let next _ := interp m in - let (pat, f) := p in - let (phyps, pconcl) := pat in - let cur _ := - let (hids, hctx, subst, cctx) := matches_goal rev phyps pconcl in - fun _ => f hids hctx subst cctx - in - Control.plus cur next - end in - Control.once (fun () => interp pats) (). - -Ltac2 multi_goal_match0 rev pats := - let rec interp m := match m with - | [] => Control.zero Match_failure - | p :: m => - let next _ := interp m in - let (pat, f) := p in - let (phyps, pconcl) := pat in - let cur _ := - let (hids, hctx, subst, cctx) := matches_goal rev phyps pconcl in - f hids hctx subst cctx - in - Control.plus cur next - end in - interp pats. - -Ltac2 one_goal_match0 rev pats := Control.once (fun _ => multi_goal_match0 rev pats). diff --git a/theories/Std.v b/theories/Std.v deleted file mode 100644 index 73b2ba02c4..0000000000 --- a/theories/Std.v +++ /dev/null @@ -1,263 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* constr), intro_pattern) -| IntroRewrite (bool) -] -with or_and_intro_pattern := [ -| IntroOrPattern (intro_pattern list list) -| IntroAndPattern (intro_pattern list) -]. - -Ltac2 Type destruction_arg := [ -| ElimOnConstr (unit -> constr_with_bindings) -| ElimOnIdent (ident) -| ElimOnAnonHyp (int) -]. - -Ltac2 Type induction_clause := { - indcl_arg : destruction_arg; - indcl_eqn : intro_pattern_naming option; - indcl_as : or_and_intro_pattern option; - indcl_in : clause option; -}. - -Ltac2 Type assertion := [ -| AssertType (intro_pattern option, constr, (unit -> unit) option) -| AssertValue (ident, constr) -]. - -Ltac2 Type repeat := [ -| Precisely (int) -| UpTo (int) -| RepeatStar -| RepeatPlus -]. - -Ltac2 Type orientation := [ LTR | RTL ]. - -Ltac2 Type rewriting := { - rew_orient : orientation option; - rew_repeat : repeat; - rew_equatn : (unit -> constr_with_bindings); -}. - -Ltac2 Type evar_flag := bool. -Ltac2 Type advanced_flag := bool. - -Ltac2 Type move_location := [ -| MoveAfter (ident) -| MoveBefore (ident) -| MoveFirst -| MoveLast -]. - -Ltac2 Type inversion_kind := [ -| SimpleInversion -| FullInversion -| FullInversionClear -]. - -(** Standard, built-in tactics. See Ltac1 for documentation. *) - -Ltac2 @ external intros : evar_flag -> intro_pattern list -> unit := "ltac2" "tac_intros". - -Ltac2 @ external apply : advanced_flag -> evar_flag -> - (unit -> constr_with_bindings) list -> (ident * (intro_pattern option)) option -> unit := "ltac2" "tac_apply". - -Ltac2 @ external elim : evar_flag -> constr_with_bindings -> constr_with_bindings option -> unit := "ltac2" "tac_elim". -Ltac2 @ external case : evar_flag -> constr_with_bindings -> unit := "ltac2" "tac_case". - -Ltac2 @ external generalize : (constr * occurrences * ident option) list -> unit := "ltac2" "tac_generalize". - -Ltac2 @ external assert : assertion -> unit := "ltac2" "tac_assert". -Ltac2 @ external enough : constr -> (unit -> unit) option option -> intro_pattern option -> unit := "ltac2" "tac_enough". - -Ltac2 @ external pose : ident option -> constr -> unit := "ltac2" "tac_pose". -Ltac2 @ external set : evar_flag -> (unit -> ident option * constr) -> clause -> unit := "ltac2" "tac_set". - -Ltac2 @ external remember : evar_flag -> ident option -> (unit -> constr) -> intro_pattern option -> clause -> unit := "ltac2" "tac_remember". - -Ltac2 @ external destruct : evar_flag -> induction_clause list -> - constr_with_bindings option -> unit := "ltac2" "tac_induction". - -Ltac2 @ external induction : evar_flag -> induction_clause list -> - constr_with_bindings option -> unit := "ltac2" "tac_induction". - -Ltac2 @ external red : clause -> unit := "ltac2" "tac_red". -Ltac2 @ external hnf : clause -> unit := "ltac2" "tac_hnf". -Ltac2 @ external simpl : red_flags -> (pattern * occurrences) option -> clause -> unit := "ltac2" "tac_simpl". -Ltac2 @ external cbv : red_flags -> clause -> unit := "ltac2" "tac_cbv". -Ltac2 @ external cbn : red_flags -> clause -> unit := "ltac2" "tac_cbn". -Ltac2 @ external lazy : red_flags -> clause -> unit := "ltac2" "tac_lazy". -Ltac2 @ external unfold : (reference * occurrences) list -> clause -> unit := "ltac2" "tac_unfold". -Ltac2 @ external fold : constr list -> clause -> unit := "ltac2" "tac_fold". -Ltac2 @ external pattern : (constr * occurrences) list -> clause -> unit := "ltac2" "tac_pattern". -Ltac2 @ external vm : (pattern * occurrences) option -> clause -> unit := "ltac2" "tac_vm". -Ltac2 @ external native : (pattern * occurrences) option -> clause -> unit := "ltac2" "tac_native". - -Ltac2 @ external eval_red : constr -> constr := "ltac2" "eval_red". -Ltac2 @ external eval_hnf : constr -> constr := "ltac2" "eval_hnf". -Ltac2 @ external eval_red : constr -> constr := "ltac2" "eval_red". -Ltac2 @ external eval_simpl : red_flags -> (pattern * occurrences) option -> constr -> constr := "ltac2" "eval_simpl". -Ltac2 @ external eval_cbv : red_flags -> constr -> constr := "ltac2" "eval_cbv". -Ltac2 @ external eval_cbn : red_flags -> constr -> constr := "ltac2" "eval_cbn". -Ltac2 @ external eval_lazy : red_flags -> constr -> constr := "ltac2" "eval_lazy". -Ltac2 @ external eval_unfold : (reference * occurrences) list -> constr -> constr := "ltac2" "eval_unfold". -Ltac2 @ external eval_fold : constr list -> constr -> constr := "ltac2" "eval_fold". -Ltac2 @ external eval_pattern : (constr * occurrences) list -> constr -> constr := "ltac2" "eval_pattern". -Ltac2 @ external eval_vm : (pattern * occurrences) option -> constr -> constr := "ltac2" "eval_vm". -Ltac2 @ external eval_native : (pattern * occurrences) option -> constr -> constr := "ltac2" "eval_native". - -Ltac2 @ external change : pattern option -> (constr array -> constr) -> clause -> unit := "ltac2" "tac_change". - -Ltac2 @ external rewrite : evar_flag -> rewriting list -> clause -> (unit -> unit) option -> unit := "ltac2" "tac_rewrite". - -Ltac2 @ external reflexivity : unit -> unit := "ltac2" "tac_reflexivity". - -Ltac2 @ external assumption : unit -> unit := "ltac2" "tac_assumption". - -Ltac2 @ external transitivity : constr -> unit := "ltac2" "tac_transitivity". - -Ltac2 @ external etransitivity : unit -> unit := "ltac2" "tac_etransitivity". - -Ltac2 @ external cut : constr -> unit := "ltac2" "tac_cut". - -Ltac2 @ external left : evar_flag -> bindings -> unit := "ltac2" "tac_left". -Ltac2 @ external right : evar_flag -> bindings -> unit := "ltac2" "tac_right". - -Ltac2 @ external constructor : evar_flag -> unit := "ltac2" "tac_constructor". -Ltac2 @ external split : evar_flag -> bindings -> unit := "ltac2" "tac_split". - -Ltac2 @ external constructor_n : evar_flag -> int -> bindings -> unit := "ltac2" "tac_constructorn". - -Ltac2 @ external intros_until : hypothesis -> unit := "ltac2" "tac_introsuntil". - -Ltac2 @ external symmetry : clause -> unit := "ltac2" "tac_symmetry". - -Ltac2 @ external rename : (ident * ident) list -> unit := "ltac2" "tac_rename". - -Ltac2 @ external revert : ident list -> unit := "ltac2" "tac_revert". - -Ltac2 @ external admit : unit -> unit := "ltac2" "tac_admit". - -Ltac2 @ external fix_ : ident option -> int -> unit := "ltac2" "tac_fix". -Ltac2 @ external cofix_ : ident option -> unit := "ltac2" "tac_cofix". - -Ltac2 @ external clear : ident list -> unit := "ltac2" "tac_clear". -Ltac2 @ external keep : ident list -> unit := "ltac2" "tac_keep". - -Ltac2 @ external clearbody : ident list -> unit := "ltac2" "tac_clearbody". - -Ltac2 @ external exact_no_check : constr -> unit := "ltac2" "tac_exactnocheck". -Ltac2 @ external vm_cast_no_check : constr -> unit := "ltac2" "tac_vmcastnocheck". -Ltac2 @ external native_cast_no_check : constr -> unit := "ltac2" "tac_nativecastnocheck". - -Ltac2 @ external inversion : inversion_kind -> destruction_arg -> intro_pattern option -> ident list option -> unit := "ltac2" "tac_inversion". - -(** coretactics *) - -Ltac2 @ external move : ident -> move_location -> unit := "ltac2" "tac_move". - -Ltac2 @ external intro : ident option -> move_location option -> unit := "ltac2" "tac_intro". - -Ltac2 @ external specialize : constr_with_bindings -> intro_pattern option -> unit := "ltac2" "tac_specialize". - -(** extratactics *) - -Ltac2 @ external discriminate : evar_flag -> destruction_arg option -> unit := "ltac2" "tac_discriminate". -Ltac2 @ external injection : evar_flag -> intro_pattern list option -> destruction_arg option -> unit := "ltac2" "tac_injection". - -Ltac2 @ external absurd : constr -> unit := "ltac2" "tac_absurd". -Ltac2 @ external contradiction : constr_with_bindings option -> unit := "ltac2" "tac_contradiction". - -Ltac2 @ external autorewrite : bool -> (unit -> unit) option -> ident list -> clause -> unit := "ltac2" "tac_autorewrite". - -Ltac2 @ external subst : ident list -> unit := "ltac2" "tac_subst". -Ltac2 @ external subst_all : unit -> unit := "ltac2" "tac_substall". - -(** auto *) - -Ltac2 Type debug := [ Off | Info | Debug ]. - -Ltac2 Type strategy := [ BFS | DFS ]. - -Ltac2 @ external trivial : debug -> (unit -> constr) list -> ident list option -> unit := "ltac2" "tac_trivial". - -Ltac2 @ external auto : debug -> int option -> (unit -> constr) list -> ident list option -> unit := "ltac2" "tac_auto". - -Ltac2 @ external new_auto : debug -> int option -> (unit -> constr) list -> ident list option -> unit := "ltac2" "tac_newauto". - -Ltac2 @ external eauto : debug -> int option -> int option -> (unit -> constr) list -> ident list option -> unit := "ltac2" "tac_eauto". - -Ltac2 @ external typeclasses_eauto : strategy option -> int option -> ident list option -> unit := "ltac2" "tac_typeclasses_eauto". - -(** firstorder *) - -Ltac2 @ external firstorder : (unit -> unit) option -> reference list -> ident list -> unit := "ltac2" "tac_firstorder". diff --git a/theories/String.v b/theories/String.v deleted file mode 100644 index 99e1dab76b..0000000000 --- a/theories/String.v +++ /dev/null @@ -1,14 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* char -> string := "ltac2" "string_make". -Ltac2 @external length : string -> int := "ltac2" "string_length". -Ltac2 @external get : string -> int -> char := "ltac2" "string_get". -Ltac2 @external set : string -> int -> char -> unit := "ltac2" "string_set". diff --git a/theories/dune b/theories/dune deleted file mode 100644 index 1fe3ba28fe..0000000000 --- a/theories/dune +++ /dev/null @@ -1,6 +0,0 @@ -(coqlib - (name Ltac2) ; This determines the -R flag - (public_name ltac2.Ltac2) - (synopsis "Ltac 2 Plugin") - (libraries ltac2.plugin)) - diff --git a/vendor/Ltac2/.gitignore b/vendor/Ltac2/.gitignore new file mode 100644 index 0000000000..00e15f8daa --- /dev/null +++ b/vendor/Ltac2/.gitignore @@ -0,0 +1,18 @@ +Makefile.coq +Makefile.coq.conf +*.glob +*.d +*.d.raw +*.vio +*.vo +*.cm* +*.annot +*.spit +*.spot +*.o +*.a +*.aux +tests/*.log +*.install +_build +.merlin diff --git a/vendor/Ltac2/.travis.yml b/vendor/Ltac2/.travis.yml new file mode 100644 index 0000000000..2628abde45 --- /dev/null +++ b/vendor/Ltac2/.travis.yml @@ -0,0 +1,40 @@ +dist: trusty +sudo: required +language: generic + +services: + - docker + +env: + global: + - NJOBS="2" + - CONTRIB_NAME="ltac2" + matrix: + - COQ_IMAGE="coqorg/coq:dev" + +install: | + # Prepare the COQ container + docker run -d -i --init --name=COQ -v ${TRAVIS_BUILD_DIR}:/home/coq/${CONTRIB_NAME} -w /home/coq/${CONTRIB_NAME} ${COQ_IMAGE} + docker exec COQ /bin/bash --login -c " + # This bash script is double-quoted to interpolate Travis CI env vars: + echo \"Build triggered by ${TRAVIS_EVENT_TYPE}\" + export PS4='+ \e[33;1m(\$0 @ line \$LINENO) \$\e[0m ' + set -ex # -e = exit on failure; -x = trace for debug + # opam update -y + # opam install -y -j ${NJOBS} coq-mathcomp-ssreflect + opam config list + opam repo list + opam list + " +script: +- echo -e "${ANSI_YELLOW}Building ${CONTRIB_NAME}...${ANSI_RESET}" && echo -en 'travis_fold:start:script\\r' +- | + docker exec COQ /bin/bash --login -c " + export PS4='+ \e[33;1m(\$0 @ line \$LINENO) \$\e[0m ' + set -ex + sudo chown -R coq:coq /home/coq/${CONTRIB_NAME} + make + make tests + " +- docker stop COQ # optional +- echo -en 'travis_fold:end:script\\r' diff --git a/vendor/Ltac2/LICENSE b/vendor/Ltac2/LICENSE new file mode 100644 index 0000000000..27950e8d20 --- /dev/null +++ b/vendor/Ltac2/LICENSE @@ -0,0 +1,458 @@ + GNU LESSER GENERAL PUBLIC LICENSE + Version 2.1, February 1999 + + Copyright (C) 1991, 1999 Free Software Foundation, Inc. + 59 Temple Place, Suite 330, Boston, MA 02111-1307 USA + Everyone is permitted to copy and distribute verbatim copies + of this license document, but changing it is not allowed. + +[This is the first released version of the Lesser GPL. It also counts + as the successor of the GNU Library Public License, version 2, hence + the version number 2.1.] + + Preamble + + The licenses for most software are designed to take away your +freedom to share and change it. By contrast, the GNU General Public +Licenses are intended to guarantee your freedom to share and change +free software--to make sure the software is free for all its users. + + This license, the Lesser General Public License, applies to some +specially designated software packages--typically libraries--of the +Free Software Foundation and other authors who decide to use it. You +can use it too, but we suggest you first think carefully about whether +this license or the ordinary General Public License is the better +strategy to use in any particular case, based on the explanations below. + + When we speak of free software, we are referring to freedom of use, +not price. Our General Public Licenses are designed to make sure that +you have the freedom to distribute copies of free software (and charge +for this service if you wish); that you receive source code or can get +it if you want it; that you can change the software and use pieces of +it in new free programs; and that you are informed that you can do +these things. + + To protect your rights, we need to make restrictions that forbid +distributors to deny you these rights or to ask you to surrender these +rights. These restrictions translate to certain responsibilities for +you if you distribute copies of the library or if you modify it. + + For example, if you distribute copies of the library, whether gratis +or for a fee, you must give the recipients all the rights that we gave +you. You must make sure that they, too, receive or can get the source +code. If you link other code with the library, you must provide +complete object files to the recipients, so that they can relink them +with the library after making changes to the library and recompiling +it. And you must show them these terms so they know their rights. + + We protect your rights with a two-step method: (1) we copyright the +library, and (2) we offer you this license, which gives you legal +permission to copy, distribute and/or modify the library. + + To protect each distributor, we want to make it very clear that +there is no warranty for the free library. Also, if the library is +modified by someone else and passed on, the recipients should know +that what they have is not the original version, so that the original +author's reputation will not be affected by problems that might be +introduced by others. + + Finally, software patents pose a constant threat to the existence of +any free program. We wish to make sure that a company cannot +effectively restrict the users of a free program by obtaining a +restrictive license from a patent holder. Therefore, we insist that +any patent license obtained for a version of the library must be +consistent with the full freedom of use specified in this license. + + Most GNU software, including some libraries, is covered by the +ordinary GNU General Public License. This license, the GNU Lesser +General Public License, applies to certain designated libraries, and +is quite different from the ordinary General Public License. We use +this license for certain libraries in order to permit linking those +libraries into non-free programs. + + When a program is linked with a library, whether statically or using +a shared library, the combination of the two is legally speaking a +combined work, a derivative of the original library. The ordinary +General Public License therefore permits such linking only if the +entire combination fits its criteria of freedom. The Lesser General +Public License permits more lax criteria for linking other code with +the library. + + We call this license the "Lesser" General Public License because it +does Less to protect the user's freedom than the ordinary General +Public License. It also provides other free software developers Less +of an advantage over competing non-free programs. These disadvantages +are the reason we use the ordinary General Public License for many +libraries. However, the Lesser license provides advantages in certain +special circumstances. + + For example, on rare occasions, there may be a special need to +encourage the widest possible use of a certain library, so that it becomes +a de-facto standard. To achieve this, non-free programs must be +allowed to use the library. A more frequent case is that a free +library does the same job as widely used non-free libraries. In this +case, there is little to gain by limiting the free library to free +software only, so we use the Lesser General Public License. + + In other cases, permission to use a particular library in non-free +programs enables a greater number of people to use a large body of +free software. For example, permission to use the GNU C Library in +non-free programs enables many more people to use the whole GNU +operating system, as well as its variant, the GNU/Linux operating +system. + + Although the Lesser General Public License is Less protective of the +users' freedom, it does ensure that the user of a program that is +linked with the Library has the freedom and the wherewithal to run +that program using a modified version of the Library. + + The precise terms and conditions for copying, distribution and +modification follow. Pay close attention to the difference between a +"work based on the library" and a "work that uses the library". The +former contains code derived from the library, whereas the latter must +be combined with the library in order to run. + + GNU LESSER GENERAL PUBLIC LICENSE + TERMS AND CONDITIONS FOR COPYING, DISTRIBUTION AND MODIFICATION + + 0. This License Agreement applies to any software library or other +program which contains a notice placed by the copyright holder or +other authorized party saying it may be distributed under the terms of +this Lesser General Public License (also called "this License"). +Each licensee is addressed as "you". + + A "library" means a collection of software functions and/or data +prepared so as to be conveniently linked with application programs +(which use some of those functions and data) to form executables. + + The "Library", below, refers to any such software library or work +which has been distributed under these terms. A "work based on the +Library" means either the Library or any derivative work under +copyright law: that is to say, a work containing the Library or a +portion of it, either verbatim or with modifications and/or translated +straightforwardly into another language. (Hereinafter, translation is +included without limitation in the term "modification".) + + "Source code" for a work means the preferred form of the work for +making modifications to it. For a library, complete source code means +all the source code for all modules it contains, plus any associated +interface definition files, plus the scripts used to control compilation +and installation of the library. + + Activities other than copying, distribution and modification are not +covered by this License; they are outside its scope. The act of +running a program using the Library is not restricted, and output from +such a program is covered only if its contents constitute a work based +on the Library (independent of the use of the Library in a tool for +writing it). Whether that is true depends on what the Library does +and what the program that uses the Library does. + + 1. You may copy and distribute verbatim copies of the Library's +complete source code as you receive it, in any medium, provided that +you conspicuously and appropriately publish on each copy an +appropriate copyright notice and disclaimer of warranty; keep intact +all the notices that refer to this License and to the absence of any +warranty; and distribute a copy of this License along with the +Library. + + You may charge a fee for the physical act of transferring a copy, +and you may at your option offer warranty protection in exchange for a +fee. + + 2. You may modify your copy or copies of the Library or any portion +of it, thus forming a work based on the Library, and copy and +distribute such modifications or work under the terms of Section 1 +above, provided that you also meet all of these conditions: + + a) The modified work must itself be a software library. + + b) You must cause the files modified to carry prominent notices + stating that you changed the files and the date of any change. + + c) You must cause the whole of the work to be licensed at no + charge to all third parties under the terms of this License. + + d) If a facility in the modified Library refers to a function or a + table of data to be supplied by an application program that uses + the facility, other than as an argument passed when the facility + is invoked, then you must make a good faith effort to ensure that, + in the event an application does not supply such function or + table, the facility still operates, and performs whatever part of + its purpose remains meaningful. + + (For example, a function in a library to compute square roots has + a purpose that is entirely well-defined independent of the + application. Therefore, Subsection 2d requires that any + application-supplied function or table used by this function must + be optional: if the application does not supply it, the square + root function must still compute square roots.) + +These requirements apply to the modified work as a whole. If +identifiable sections of that work are not derived from the Library, +and can be reasonably considered independent and separate works in +themselves, then this License, and its terms, do not apply to those +sections when you distribute them as separate works. But when you +distribute the same sections as part of a whole which is a work based +on the Library, the distribution of the whole must be on the terms of +this License, whose permissions for other licensees extend to the +entire whole, and thus to each and every part regardless of who wrote +it. + +Thus, it is not the intent of this section to claim rights or contest +your rights to work written entirely by you; rather, the intent is to +exercise the right to control the distribution of derivative or +collective works based on the Library. + +In addition, mere aggregation of another work not based on the Library +with the Library (or with a work based on the Library) on a volume of +a storage or distribution medium does not bring the other work under +the scope of this License. + + 3. You may opt to apply the terms of the ordinary GNU General Public +License instead of this License to a given copy of the Library. To do +this, you must alter all the notices that refer to this License, so +that they refer to the ordinary GNU General Public License, version 2, +instead of to this License. (If a newer version than version 2 of the +ordinary GNU General Public License has appeared, then you can specify +that version instead if you wish.) Do not make any other change in +these notices. + + Once this change is made in a given copy, it is irreversible for +that copy, so the ordinary GNU General Public License applies to all +subsequent copies and derivative works made from that copy. + + This option is useful when you wish to copy part of the code of +the Library into a program that is not a library. + + 4. You may copy and distribute the Library (or a portion or +derivative of it, under Section 2) in object code or executable form +under the terms of Sections 1 and 2 above provided that you accompany +it with the complete corresponding machine-readable source code, which +must be distributed under the terms of Sections 1 and 2 above on a +medium customarily used for software interchange. + + If distribution of object code is made by offering access to copy +from a designated place, then offering equivalent access to copy the +source code from the same place satisfies the requirement to +distribute the source code, even though third parties are not +compelled to copy the source along with the object code. + + 5. A program that contains no derivative of any portion of the +Library, but is designed to work with the Library by being compiled or +linked with it, is called a "work that uses the Library". Such a +work, in isolation, is not a derivative work of the Library, and +therefore falls outside the scope of this License. + + However, linking a "work that uses the Library" with the Library +creates an executable that is a derivative of the Library (because it +contains portions of the Library), rather than a "work that uses the +library". The executable is therefore covered by this License. +Section 6 states terms for distribution of such executables. + + When a "work that uses the Library" uses material from a header file +that is part of the Library, the object code for the work may be a +derivative work of the Library even though the source code is not. +Whether this is true is especially significant if the work can be +linked without the Library, or if the work is itself a library. The +threshold for this to be true is not precisely defined by law. + + If such an object file uses only numerical parameters, data +structure layouts and accessors, and small macros and small inline +functions (ten lines or less in length), then the use of the object +file is unrestricted, regardless of whether it is legally a derivative +work. (Executables containing this object code plus portions of the +Library will still fall under Section 6.) + + Otherwise, if the work is a derivative of the Library, you may +distribute the object code for the work under the terms of Section 6. +Any executables containing that work also fall under Section 6, +whether or not they are linked directly with the Library itself. + + 6. As an exception to the Sections above, you may also combine or +link a "work that uses the Library" with the Library to produce a +work containing portions of the Library, and distribute that work +under terms of your choice, provided that the terms permit +modification of the work for the customer's own use and reverse +engineering for debugging such modifications. + + You must give prominent notice with each copy of the work that the +Library is used in it and that the Library and its use are covered by +this License. You must supply a copy of this License. If the work +during execution displays copyright notices, you must include the +copyright notice for the Library among them, as well as a reference +directing the user to the copy of this License. Also, you must do one +of these things: + + a) Accompany the work with the complete corresponding + machine-readable source code for the Library including whatever + changes were used in the work (which must be distributed under + Sections 1 and 2 above); and, if the work is an executable linked + with the Library, with the complete machine-readable "work that + uses the Library", as object code and/or source code, so that the + user can modify the Library and then relink to produce a modified + executable containing the modified Library. (It is understood + that the user who changes the contents of definitions files in the + Library will not necessarily be able to recompile the application + to use the modified definitions.) + + b) Use a suitable shared library mechanism for linking with the + Library. A suitable mechanism is one that (1) uses at run time a + copy of the library already present on the user's computer system, + rather than copying library functions into the executable, and (2) + will operate properly with a modified version of the library, if + the user installs one, as long as the modified version is + interface-compatible with the version that the work was made with. + + c) Accompany the work with a written offer, valid for at + least three years, to give the same user the materials + specified in Subsection 6a, above, for a charge no more + than the cost of performing this distribution. + + d) If distribution of the work is made by offering access to copy + from a designated place, offer equivalent access to copy the above + specified materials from the same place. + + e) Verify that the user has already received a copy of these + materials or that you have already sent this user a copy. + + For an executable, the required form of the "work that uses the +Library" must include any data and utility programs needed for +reproducing the executable from it. However, as a special exception, +the materials to be distributed need not include anything that is +normally distributed (in either source or binary form) with the major +components (compiler, kernel, and so on) of the operating system on +which the executable runs, unless that component itself accompanies +the executable. + + It may happen that this requirement contradicts the license +restrictions of other proprietary libraries that do not normally +accompany the operating system. Such a contradiction means you cannot +use both them and the Library together in an executable that you +distribute. + + 7. You may place library facilities that are a work based on the +Library side-by-side in a single library together with other library +facilities not covered by this License, and distribute such a combined +library, provided that the separate distribution of the work based on +the Library and of the other library facilities is otherwise +permitted, and provided that you do these two things: + + a) Accompany the combined library with a copy of the same work + based on the Library, uncombined with any other library + facilities. This must be distributed under the terms of the + Sections above. + + b) Give prominent notice with the combined library of the fact + that part of it is a work based on the Library, and explaining + where to find the accompanying uncombined form of the same work. + + 8. You may not copy, modify, sublicense, link with, or distribute +the Library except as expressly provided under this License. Any +attempt otherwise to copy, modify, sublicense, link with, or +distribute the Library is void, and will automatically terminate your +rights under this License. However, parties who have received copies, +or rights, from you under this License will not have their licenses +terminated so long as such parties remain in full compliance. + + 9. You are not required to accept this License, since you have not +signed it. However, nothing else grants you permission to modify or +distribute the Library or its derivative works. These actions are +prohibited by law if you do not accept this License. Therefore, by +modifying or distributing the Library (or any work based on the +Library), you indicate your acceptance of this License to do so, and +all its terms and conditions for copying, distributing or modifying +the Library or works based on it. + + 10. Each time you redistribute the Library (or any work based on the +Library), the recipient automatically receives a license from the +original licensor to copy, distribute, link with or modify the Library +subject to these terms and conditions. You may not impose any further +restrictions on the recipients' exercise of the rights granted herein. +You are not responsible for enforcing compliance by third parties with +this License. + + 11. If, as a consequence of a court judgment or allegation of patent +infringement or for any other reason (not limited to patent issues), +conditions are imposed on you (whether by court order, agreement or +otherwise) that contradict the conditions of this License, they do not +excuse you from the conditions of this License. If you cannot +distribute so as to satisfy simultaneously your obligations under this +License and any other pertinent obligations, then as a consequence you +may not distribute the Library at all. For example, if a patent +license would not permit royalty-free redistribution of the Library by +all those who receive copies directly or indirectly through you, then +the only way you could satisfy both it and this License would be to +refrain entirely from distribution of the Library. + +If any portion of this section is held invalid or unenforceable under any +particular circumstance, the balance of the section is intended to apply, +and the section as a whole is intended to apply in other circumstances. + +It is not the purpose of this section to induce you to infringe any +patents or other property right claims or to contest validity of any +such claims; this section has the sole purpose of protecting the +integrity of the free software distribution system which is +implemented by public license practices. Many people have made +generous contributions to the wide range of software distributed +through that system in reliance on consistent application of that +system; it is up to the author/donor to decide if he or she is willing +to distribute software through any other system and a licensee cannot +impose that choice. + +This section is intended to make thoroughly clear what is believed to +be a consequence of the rest of this License. + + 12. If the distribution and/or use of the Library is restricted in +certain countries either by patents or by copyrighted interfaces, the +original copyright holder who places the Library under this License may add +an explicit geographical distribution limitation excluding those countries, +so that distribution is permitted only in or among countries not thus +excluded. In such case, this License incorporates the limitation as if +written in the body of this License. + + 13. The Free Software Foundation may publish revised and/or new +versions of the Lesser General Public License from time to time. +Such new versions will be similar in spirit to the present version, +but may differ in detail to address new problems or concerns. + +Each version is given a distinguishing version number. If the Library +specifies a version number of this License which applies to it and +"any later version", you have the option of following the terms and +conditions either of that version or of any later version published by +the Free Software Foundation. If the Library does not specify a +license version number, you may choose any version ever published by +the Free Software Foundation. + + 14. If you wish to incorporate parts of the Library into other free +programs whose distribution conditions are incompatible with these, +write to the author to ask for permission. For software which is +copyrighted by the Free Software Foundation, write to the Free +Software Foundation; we sometimes make exceptions for this. Our +decision will be guided by the two goals of preserving the free status +of all derivatives of our free software and of promoting the sharing +and reuse of software generally. + + NO WARRANTY + + 15. BECAUSE THE LIBRARY IS LICENSED FREE OF CHARGE, THERE IS NO +WARRANTY FOR THE LIBRARY, TO THE EXTENT PERMITTED BY APPLICABLE LAW. +EXCEPT WHEN OTHERWISE STATED IN WRITING THE COPYRIGHT HOLDERS AND/OR +OTHER PARTIES PROVIDE THE LIBRARY "AS IS" WITHOUT WARRANTY OF ANY +KIND, EITHER EXPRESSED OR IMPLIED, INCLUDING, BUT NOT LIMITED TO, THE +IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR +PURPOSE. THE ENTIRE RISK AS TO THE QUALITY AND PERFORMANCE OF THE +LIBRARY IS WITH YOU. SHOULD THE LIBRARY PROVE DEFECTIVE, YOU ASSUME +THE COST OF ALL NECESSARY SERVICING, REPAIR OR CORRECTION. + + 16. IN NO EVENT UNLESS REQUIRED BY APPLICABLE LAW OR AGREED TO IN +WRITING WILL ANY COPYRIGHT HOLDER, OR ANY OTHER PARTY WHO MAY MODIFY +AND/OR REDISTRIBUTE THE LIBRARY AS PERMITTED ABOVE, BE LIABLE TO YOU +FOR DAMAGES, INCLUDING ANY GENERAL, SPECIAL, INCIDENTAL OR +CONSEQUENTIAL DAMAGES ARISING OUT OF THE USE OR INABILITY TO USE THE +LIBRARY (INCLUDING BUT NOT LIMITED TO LOSS OF DATA OR DATA BEING +RENDERED INACCURATE OR LOSSES SUSTAINED BY YOU OR THIRD PARTIES OR A +FAILURE OF THE LIBRARY TO OPERATE WITH ANY OTHER SOFTWARE), EVEN IF +SUCH HOLDER OR OTHER PARTY HAS BEEN ADVISED OF THE POSSIBILITY OF SUCH +DAMAGES. + + END OF TERMS AND CONDITIONS diff --git a/vendor/Ltac2/Makefile b/vendor/Ltac2/Makefile new file mode 100644 index 0000000000..e0e197650d --- /dev/null +++ b/vendor/Ltac2/Makefile @@ -0,0 +1,14 @@ +ifeq "$(COQBIN)" "" + COQBIN=$(dir $(shell which coqtop))/ +endif + +%: Makefile.coq + +Makefile.coq: _CoqProject + $(COQBIN)coq_makefile -f _CoqProject -o Makefile.coq + +tests: all + @$(MAKE) -C tests -s clean + @$(MAKE) -C tests -s all + +-include Makefile.coq diff --git a/vendor/Ltac2/README.md b/vendor/Ltac2/README.md new file mode 100644 index 0000000000..d49dd88076 --- /dev/null +++ b/vendor/Ltac2/README.md @@ -0,0 +1,25 @@ +[![Build Status](https://travis-ci.org/ppedrot/ltac2.svg?branch=master)](https://travis-ci.org/ppedrot/ltac2) + +Overview +======== + +This is a standalone version of the Ltac2 plugin. Ltac2 is an attempt at +providing the Coq users with a tactic language that is more robust and more +expressive than the venerable Ltac language. + +Status +======== + +It is mostly a toy to experiment for now, and the implementation is quite +bug-ridden. Don't mistake this for a final product! + +Installation +============ + +This should compile with Coq master, assuming the `COQBIN` variable is set +correctly. Standard procedures for `coq_makefile`-generated plugins apply. + +Demo +==== + +Horrible test-files are provided in the `tests` folder. Not for kids. diff --git a/vendor/Ltac2/_CoqProject b/vendor/Ltac2/_CoqProject new file mode 100644 index 0000000000..dda5a8001a --- /dev/null +++ b/vendor/Ltac2/_CoqProject @@ -0,0 +1,51 @@ +-R theories/ Ltac2 +-I src/ + +src/tac2dyn.ml +src/tac2dyn.mli +src/tac2expr.mli +src/tac2types.mli +src/tac2env.ml +src/tac2env.mli +src/tac2print.ml +src/tac2print.mli +src/tac2intern.ml +src/tac2intern.mli +src/tac2interp.ml +src/tac2interp.mli +src/tac2entries.ml +src/tac2entries.mli +src/tac2ffi.ml +src/tac2ffi.mli +src/tac2qexpr.mli +src/tac2quote.ml +src/tac2quote.mli +src/tac2match.ml +src/tac2match.mli +src/tac2core.ml +src/tac2core.mli +src/tac2extffi.ml +src/tac2extffi.mli +src/tac2tactics.ml +src/tac2tactics.mli +src/tac2stdlib.ml +src/tac2stdlib.mli +src/g_ltac2.mlg +src/ltac2_plugin.mlpack + +theories/Init.v +theories/Int.v +theories/Char.v +theories/String.v +theories/Ident.v +theories/Array.v +theories/Control.v +theories/Message.v +theories/Constr.v +theories/Pattern.v +theories/Fresh.v +theories/Std.v +theories/Env.v +theories/Notations.v +theories/Ltac1.v +theories/Ltac2.v diff --git a/vendor/Ltac2/doc/ltac2.md b/vendor/Ltac2/doc/ltac2.md new file mode 100644 index 0000000000..b217cb08e6 --- /dev/null +++ b/vendor/Ltac2/doc/ltac2.md @@ -0,0 +1,1036 @@ +# Summary + +The Ltac tactic language is probably one of the ingredients of the success of +Coq, yet it is at the same time its Achilles' heel. Indeed, Ltac: + +- has nothing like intended semantics +- is very non-uniform due to organic growth +- lacks expressivity and requires programming-by-hacking +- is slow +- is error-prone and fragile +- has an intricate implementation + +This has a lot of terrible consequences, most notably the fact that it is never +clear whether some observed behaviour is a bug or a proper one. + +Following the need of users that start developing huge projects relying +critically on Ltac, we believe that we should offer a proper modern language +that features at least the following: + +- at least informal, predictable semantics +- a typing system +- standard programming facilities (i.e. datatypes) + +This document describes the implementation of such a language. The +implementation of Ltac as of Coq 8.7 will be referred to as Ltac1. + +# Contents + + +**Table of Contents** + +- [Summary](#summary) +- [Contents](#contents) +- [General design](#general-design) +- [ML component](#ml-component) + - [Overview](#overview) + - [Type Syntax](#type-syntax) + - [Type declarations](#type-declarations) + - [Term Syntax](#term-syntax) + - [Ltac Definitions](#ltac-definitions) + - [Reduction](#reduction) + - [Typing](#typing) + - [Effects](#effects) + - [Standard IO](#standard-io) + - [Fatal errors](#fatal-errors) + - [Backtrack](#backtrack) + - [Goals](#goals) +- [Meta-programming](#meta-programming) + - [Overview](#overview-1) + - [Generic Syntax for Quotations](#generic-syntax-for-quotations) + - [Built-in quotations](#built-in-quotations) + - [Strict vs. non-strict mode](#strict-vs-non-strict-mode) + - [Term Antiquotations](#term-antiquotations) + - [Syntax](#syntax) + - [Semantics](#semantics) + - [Static semantics](#static-semantics) + - [Dynamic semantics](#dynamic-semantics) + - [Trivial Term Antiquotations](#trivial-term-antiquotations) + - [Match over terms](#match-over-terms) + - [Match over goals](#match-over-goals) +- [Notations](#notations) + - [Scopes](#scopes) + - [Notations](#notations-1) + - [Abbreviations](#abbreviations) +- [Evaluation](#evaluation) +- [Debug](#debug) +- [Compatibility layer with Ltac1](#compatibility-layer-with-ltac1) + - [Ltac1 from Ltac2](#ltac1-from-ltac2) + - [Ltac2 from Ltac1](#ltac2-from-ltac1) +- [Transition from Ltac1](#transition-from-ltac1) + - [Syntax changes](#syntax-changes) + - [Tactic delay](#tactic-delay) + - [Variable binding](#variable-binding) + - [In Ltac expressions](#in-ltac-expressions) + - [In quotations](#in-quotations) + - [Exception catching](#exception-catching) +- [TODO](#todo) + + + + +# General design + +There are various alternatives to Ltac1, such that Mtac or Rtac for instance. +While those alternatives can be quite distinct from Ltac1, we designed +Ltac2 to be closest as reasonably possible to Ltac1, while fixing the +aforementioned defects. + +In particular, Ltac2 is: +- a member of the ML family of languages, i.e. + * a call-by-value functional language + * with effects + * together with Hindley-Milner type system +- a language featuring meta-programming facilities for the manipulation of + Coq-side terms +- a language featuring notation facilities to help writing palatable scripts + +We describe more in details each point in the remainder of this document. + +# ML component + +## Overview + +Ltac2 is a member of the ML family of languages, in the sense that it is an +effectful call-by-value functional language, with static typing à la +Hindley-Milner. It is commonly accepted that ML constitutes a sweet spot in PL +design, as it is relatively expressive while not being either too lax +(contrarily to dynamic typing) nor too strict (contrarily to say, dependent +types). + +The main goal of Ltac2 is to serve as a meta-language for Coq. As such, it +naturally fits in the ML lineage, just as the historical ML was designed as +the tactic language for the LCF prover. It can also be seen as a general-purpose +language, by simply forgetting about the Coq-specific features. + +Sticking to a standard ML type system can be considered somewhat weak for a +meta-language designed to manipulate Coq terms. In particular, there is no +way to statically guarantee that a Coq term resulting from an Ltac2 +computation will be well-typed. This is actually a design choice, motivated +by retro-compatibility with Ltac1. Instead, well-typedness is deferred to +dynamic checks, allowing many primitive functions to fail whenever they are +provided with an ill-typed term. + +The language is naturally effectful as it manipulates the global state of the +proof engine. This allows to think of proof-modifying primitives as effects +in a straightforward way. Semantically, proof manipulation lives in a monad, +which allows to ensure that Ltac2 satisfies the same equations as a generic ML +with unspecified effects would do, e.g. function reduction is substitution +by a value. + +## Type Syntax + +At the level of terms, we simply elaborate on Ltac1 syntax, which is quite +close to e.g. the one of OCaml. Types follow the simply-typed syntax of OCaml. + +``` +TYPE := +| "(" TYPE₀ "," ... "," TYPEₙ ")" TYPECONST +| "(" TYPE₀ "*" ... "*" TYPEₙ ")" +| TYPE₁ "->" TYPE₂ +| TYPEVAR + +TYPECONST := ( MODPATH "." )* LIDENT + +TYPEVAR := "'" LIDENT + +TYPEPARAMS := "(" TYPEVAR₀ "," ... "," TYPEVARₙ ")" +``` + +The set of base types can be extended thanks to the usual ML type +declarations such as algebraic datatypes and records. + +Built-in types include: +- `int`, machine integers (size not specified, in practice inherited from OCaml) +- `string`, mutable strings +- `'a array`, mutable arrays +- `exn`, exceptions +- `constr`, kernel-side terms +- `pattern`, term patterns +- `ident`, well-formed identifiers + +## Type declarations + +One can define new types by the following commands. + +``` +VERNAC ::= +| "Ltac2" "Type" TYPEPARAMS LIDENT +| "Ltac2" "Type" RECFLAG TYPEPARAMS LIDENT ":=" TYPEDEF + +RECFLAG := ( "rec" ) +``` + +The first command defines an abstract type. It has no use for the end user and +is dedicated to types representing data coming from the OCaml world. + +The second command defines a type with a manifest. There are four possible +kinds of such definitions: alias, variant, record and open variant types. + +``` +TYPEDEF := +| TYPE +| "[" CONSTRUCTORDEF₀ "|" ... "|" CONSTRUCTORDEFₙ "]" +| "{" FIELDDEF₀ ";" ... ";" FIELDDEFₙ "}" +| "[" ".." "]" + +CONSTRUCTORDEF := +| IDENT ( "(" TYPE₀ "," ... "," TYPE₀ ")" ) + +FIELDDEF := +| MUTFLAG IDENT ":" TYPE + +MUTFLAG := ( "mutable" ) +``` + +Aliases are just a name for a given type expression and are transparently +unfoldable to it. They cannot be recursive. + +Variants are sum types defined by constructors and eliminated by +pattern-matching. They can be recursive, but the `RECFLAG` must be explicitly +set. Pattern-maching must be exhaustive. + +Records are product types with named fields and eliminated by projection. +Likewise they can be recursive if the `RECFLAG` is set. + +Open variants are a special kind of variant types whose constructors are not +statically defined, but can instead be extended dynamically. A typical example +is the standard `exn` type. Pattern-matching must always include a catch-all +clause. They can be extended by the following command. + +``` +VERNAC ::= +| "Ltac2" "Type" TYPEPARAMS QUALID ":=" "[" CONSTRUCTORDEF "]" +``` + +## Term Syntax + +The syntax of the functional fragment is very close to the one of Ltac1, except +that it adds a true pattern-matching feature, as well as a few standard +constructions from ML. + +``` +VAR := LIDENT + +QUALID := ( MODPATH "." )* LIDENT + +CONSTRUCTOR := UIDENT + +TERM := +| QUALID +| CONSTRUCTOR TERM₀ ... TERMₙ +| TERM TERM₀ ... TERMₙ +| "fun" VAR "=>" TERM +| "let" VAR ":=" TERM "in" TERM +| "let" "rec" VAR ":=" TERM "in" TERM +| "match" TERM "with" BRANCH* "end" +| INT +| STRING +| "[|" TERM₀ ";" ... ";" TERMₙ "|]" +| "(" TERM₀ "," ... "," TERMₙ ")" +| "{" FIELD+ "}" +| TERM "." "(" QUALID ")" +| TERM₁ "." "(" QUALID ")" ":=" TERM₂ +| "["; TERM₀ ";" ... ";" TERMₙ "]" +| TERM₁ "::" TERM₂ +| ... + +BRANCH := +| PATTERN "=>" TERM + +PATTERN := +| VAR +| "_" +| "(" PATTERN₀ "," ... "," PATTERNₙ ")" +| CONSTRUCTOR PATTERN₀ ... PATTERNₙ +| "[" "]" +| PATTERN₁ "::" PATTERN₂ + +FIELD := +| QUALID ":=" TERM + +``` + +In practice, there is some additional syntactic sugar that allows e.g. to +bind a variable and match on it at the same time, in the usual ML style. + +There is a dedicated syntax for list and array litterals. + +Limitations: for now, deep pattern matching is not implemented yet. + +## Ltac Definitions + +One can define a new global Ltac2 value using the following syntax. +``` +VERNAC ::= +| "Ltac2" MUTFLAG RECFLAG LIDENT ":=" TERM +``` + +For semantic reasons, the body of the Ltac2 definition must be a syntactical +value, i.e. a function, a constant or a pure constructor recursively applied to +values. + +If the `RECFLAG` is set, the tactic is expanded into a recursive binding. + +If the `MUTFLAG` is set, the definition can be redefined at a later stage. This +can be performed through the following command. + +``` +VERNAC ::= +| "Ltac2" "Set" QUALID ":=" TERM +``` + +Mutable definitions act like dynamic binding, i.e. at runtime, the last defined +value for this entry is chosen. This is useful for global flags and the like. + +## Reduction + +We use the usual ML call-by-value reduction, with an otherwise unspecified +evaluation order. This is a design choice making it compatible with OCaml, +if ever we implement native compilation. The expected equations are as follows. +``` +(fun x => t) V ≡ t{x := V} (βv) + +let x := V in t ≡ t{x := V} (let) + +match C V₀ ... Vₙ with ... | C x₀ ... xₙ => t | ... end ≡ t {xᵢ := Vᵢ} (ι) + +(t any term, V values, C constructor) +``` + +Note that call-by-value reduction is already a departure from Ltac1 which uses +heuristics to decide when evaluating an expression. For instance, the following +expressions do not evaluate the same way in Ltac1. + +``` +foo (idtac; let x := 0 in bar) + +foo (let x := 0 in bar) +``` + +Instead of relying on the `idtac` hack, we would now require an explicit thunk +not to compute the argument, and `foo` would have e.g. type +`(unit -> unit) -> unit`. + +``` +foo (fun () => let x := 0 in bar) +``` + +## Typing + +Typing is strict and follows Hindley-Milner system. We will not implement the +current hackish subtyping semantics, and one will have to resort to conversion +functions. See notations though to make things more palatable. + +In this setting, all usual argument-free tactics have type `unit -> unit`, but +one can return as well a value of type `t` thanks to terms of type `unit -> t`, +or take additional arguments. + +## Effects + +Regarding effects, nothing involved here, except that instead of using the +standard IO monad as the ambient effectful world, Ltac2 is going to use the +tactic monad. + +Note that the order of evaluation of application is *not* specified and is +implementation-dependent, as in OCaml. + +We recall that the `Proofview.tactic` monad is essentially a IO monad together +with backtracking state representing the proof state. + +Intuitively a thunk of type `unit -> 'a` can do the following: +- It can perform non-backtracking IO like printing and setting mutable variables +- It can fail in a non-recoverable way +- It can use first-class backtrack. The proper way to figure that is that we + morally have the following isomorphism: + `(unit -> 'a) ~ (unit -> exn + ('a * (exn -> 'a)))` i.e. thunks can produce a + lazy list of results where each tail is waiting for a continuation exception. +- It can access a backtracking proof state, made out amongst other things of + the current evar assignation and the list of goals under focus. + +We describe more thoroughly the various effects existing in Ltac2 hereafter. + +### Standard IO + +The Ltac2 language features non-backtracking IO, notably mutable data and +printing operations. + +Mutable fields of records can be modified using the set syntax. Likewise, +built-in types like `string` and `array` feature imperative assignment. See +modules `String` and `Array` respectively. + +A few printing primitives are provided in the `Message` module, allowing to +display information to the user. + +### Fatal errors + +The Ltac2 language provides non-backtracking exceptions through the +following primitive in module `Control`. + +``` +val throw : exn -> 'a +``` + +Contrarily to backtracking exceptions from the next section, this kind of error +is never caught by backtracking primitives, that is, throwing an exception +destroys the stack. This is materialized by the following equation, where `E` +is an evaluation context. + +``` +E[throw e] ≡ throw e + +(e value) +``` + +There is currently no way to catch such an exception and it is a design choice. +There might be at some future point a way to catch it in a brutal way, +destroying all backtrack and return values. + +### Backtrack + +In Ltac2, we have the following backtracking primitives, defined in the +`Control` module. + +``` +Ltac2 Type 'a result := [ Val ('a) | Err (exn) ]. + +val zero : exn -> 'a +val plus : (unit -> 'a) -> (exn -> 'a) -> 'a +val case : (unit -> 'a) -> ('a * (exn -> 'a)) result +``` + +If one sees thunks as lazy lists, then `zero` is the empty list and `plus` is +list concatenation, while `case` is pattern-matching. + +The backtracking is first-class, i.e. one can write +`plus (fun () => "x") (fun _ => "y") : string` producing a backtracking string. + +These operations are expected to satisfy a few equations, most notably that they +form a monoid compatible with sequentialization. + +``` +plus t zero ≡ t () +plus (fun () => zero e) f ≡ f e +plus (plus t f) g ≡ plus t (fun e => plus (f e) g) + +case (fun () => zero e) ≡ Err e +case (fun () => plus (fun () => t) f) ≡ Val t f + +let x := zero e in u ≡ zero e +let x := plus t f in u ≡ plus (fun () => let x := t in u) (fun e => let x := f e in u) + +(t, u, f, g, e values) +``` + +### Goals + +A goal is given by the data of its conclusion and hypotheses, i.e. it can be +represented as `[Γ ⊢ A]`. + +The tactic monad naturally operates over the whole proofview, which may +represent several goals, including none. Thus, there is no such thing as +*the current goal*. Goals are naturally ordered, though. + +It is natural to do the same in Ltac2, but we must provide a way to get access +to a given goal. This is the role of the `enter` primitive, that applies a +tactic to each currently focused goal in turn. + +``` +val enter : (unit -> unit) -> unit +``` + +It is guaranteed that when evaluating `enter f`, `f` is called with exactly one +goal under focus. Note that `f` may be called several times, or never, depending +on the number of goals under focus before the call to `enter`. + +Accessing the goal data is then implicit in the Ltac2 primitives, and may panic +if the invariants are not respected. The two essential functions for observing +goals are given below. + +``` +val hyp : ident -> constr +val goal : unit -> constr +``` + +The two above functions panic if there is not exactly one goal under focus. +In addition, `hyp` may also fail if there is no hypothesis with the +corresponding name. + +# Meta-programming + +## Overview + +One of the horrendous implementation issues of Ltac is the fact that it is +never clear whether an object refers to the object world or the meta-world. +This is an incredible source of slowness, as the interpretation must be +aware of bound variables and must use heuristics to decide whether a variable +is a proper one or referring to something in the Ltac context. + +Likewise, in Ltac1, constr parsing is implicit, so that `foo 0` is +not `foo` applied to the Ltac integer expression `0` (Ltac does have a +non-first-class notion of integers), but rather the Coq term `Datatypes.O`. + +We should stop doing that! We need to mark when quoting and unquoting, although +we need to do that in a short and elegant way so as not to be too cumbersome +to the user. + +## Generic Syntax for Quotations + +In general, quotations can be introduced in term by the following syntax, where +`QUOTENTRY` is some parsing entry. +``` +TERM ::= +| QUOTNAME ":" "(" QUOTENTRY ")" + +QUOTNAME := IDENT +``` + +### Built-in quotations + +The current implementation recognizes the following built-in quotations: +- "ident", which parses identifiers (type `Init.ident`). +- "constr", which parses Coq terms and produces an-evar free term at runtime + (type `Init.constr`). +- "open_constr", which parses Coq terms and produces a term potentially with + holes at runtime (type `Init.constr` as well). +- "pattern", which parses Coq patterns and produces a pattern used for term + matching (type `Init.pattern`). +- "reference", which parses either a `QUALID` or `"&" IDENT`. Qualified names + are globalized at internalization into the corresponding global reference, + while `&id` is turned into `Std.VarRef id`. This produces at runtime a + `Std.reference`. + +The following syntactic sugar is provided for two common cases. +- `@id` is the same as ident:(id) +- `'t` is the same as open_constr:(t) + +### Strict vs. non-strict mode + +Depending on the context, quotations producing terms (i.e. `constr` or +`open_constr`) are not internalized in the same way. There are two possible +modes, respectively called the *strict* and the *non-strict* mode. + +- In strict mode, all simple identifiers appearing in a term quotation are +required to be resolvable statically. That is, they must be the short name of +a declaration which is defined globally, excluding section variables and +hypotheses. If this doesn't hold, internalization will fail. To work around +this error, one has to specifically use the `&` notation. +- In non-strict mode, any simple identifier appearing in a term quotation which +is not bound in the global context is turned into a dynamic reference to a +hypothesis. That is to say, internalization will succeed, but the evaluation +of the term at runtime will fail if there is no such variable in the dynamic +context. + +Strict mode is enforced by default, e.g. for all Ltac2 definitions. Non-strict +mode is only set when evaluating Ltac2 snippets in interactive proof mode. The +rationale is that it is cumbersome to explicitly add `&` interactively, while it +is expected that global tactics enforce more invariants on their code. + +## Term Antiquotations + +### Syntax + +One can also insert Ltac2 code into Coq term, similarly to what was possible in +Ltac1. + +``` +COQCONSTR ::= +| "ltac2" ":" "(" TERM ")" +``` + +Antiquoted terms are expected to have type `unit`, as they are only evaluated +for their side-effects. + +### Semantics + +Interpretation of a quoted Coq term is done in two phases, internalization and +evaluation. + +- Internalization is part of the static semantics, i.e. it is done at Ltac2 + typing time. +- Evaluation is part of the dynamic semantics, i.e. it is done when + a term gets effectively computed by Ltac2. + +Remark that typing of Coq terms is a *dynamic* process occuring at Ltac2 +evaluation time, and not at Ltac2 typing time. + +#### Static semantics + +During internalization, Coq variables are resolved and antiquotations are +type-checked as Ltac2 terms, effectively producing a `glob_constr` in Coq +implementation terminology. Note that although it went through the +type-checking of **Ltac2**, the resulting term has not been fully computed and +is potentially ill-typed as a runtime **Coq** term. + +``` +Ltac2 Definition myconstr () := constr:(nat -> 0). +// Valid with type `unit -> constr`, but will fail at runtime. +``` + +Term antiquotations are type-checked in the enclosing Ltac2 typing context +of the corresponding term expression. For instance, the following will +type-check. + +``` +let x := '0 in constr:(1 + ltac2:(exact x)) +// type constr +``` + +Beware that the typing environment of typing of antiquotations is **not** +expanded by the Coq binders from the term. Namely, it means that the following +Ltac2 expression will **not** type-check. +``` +constr:(fun x : nat => ltac2:(exact x)) +// Error: Unbound variable 'x' +``` + +There is a simple reason for that, which is that the following expression would +not make sense in general. +``` +constr:(fun x : nat => ltac2:(clear @x; exact x)) +``` +Indeed, a hypothesis can suddenly disappear from the runtime context if some +other tactic pulls the rug from under you. + +Rather, the tactic writer has to resort to the **dynamic** goal environment, +and must write instead explicitly that she is accessing a hypothesis, typically +as follows. +``` +constr:(fun x : nat => ltac2:(exact (hyp @x))) +``` + +This pattern is so common that we provide dedicated Ltac2 and Coq term notations +for it. + +- `&x` as an Ltac2 expression expands to `hyp @x`. +- `&x` as a Coq constr expression expands to + `ltac2:(Control.refine (fun () => hyp @x))`. + +#### Dynamic semantics + +During evaluation, a quoted term is fully evaluated to a kernel term, and is +in particular type-checked in the current environment. + +Evaluation of a quoted term goes as follows. +- The quoted term is first evaluated by the pretyper. +- Antiquotations are then evaluated in a context where there is exactly one goal +under focus, with the hypotheses coming from the current environment extended +with the bound variables of the term, and the resulting term is fed into the +quoted term. + +Relative orders of evaluation of antiquotations and quoted term are not +specified. + +For instance, in the following example, `tac` will be evaluated in a context +with exactly one goal under focus, whose last hypothesis is `H : nat`. The +whole expression will thus evaluate to the term `fun H : nat => nat`. +``` +let tac () := hyp @H in constr:(fun H : nat => ltac2:(tac ())) +``` + +Many standard tactics perform type-checking of their argument before going +further. It is your duty to ensure that terms are well-typed when calling +such tactics. Failure to do so will result in non-recoverable exceptions. + +## Trivial Term Antiquotations + +It is possible to refer to a variable of type `constr` in the Ltac2 environment +through a specific syntax consistent with the antiquotations presented in +the notation section. + +``` +COQCONSTR ::= +| "$" LIDENT +``` + +In a Coq term, writing `$x` is semantically equivalent to +`ltac2:(Control.refine (fun () => x))`, up to re-typechecking. It allows to +insert in a concise way an Ltac2 variable of type `constr` into a Coq term. + +## Match over terms + +Ltac2 features a construction similar to Ltac1 `match` over terms, although +in a less hard-wired way. + +``` +TERM ::= +| "match!" TERM "with" CONSTR-MATCHING* "end" +| "lazy_match!" TERM "with" CONSTR-MATCHING* "end" +| "multi_match!" TERM "with" CONSTR-MATCHING*"end" + +CONSTR-MATCHING := +| "|" CONSTR-PATTERN "=>" TERM + +CONSTR-PATTERN := +| CONSTR +| "context" LIDENT? "[" CONSTR "]" +``` + +This construction is not primitive and is desugared at parsing time into +calls to term matching functions from the `Pattern` module. Internally, it is +implemented thanks to a specific scope accepting the `CONSTR-MATCHING` syntax. + +Variables from the `CONSTR-PATTERN` are statically bound in the body of the branch, to +values of type `constr` for the variables from the `CONSTR` pattern and to a +value of type `Pattern.context` for the variable `LIDENT`. + +Note that contrarily to Ltac, only lowercase identifiers are valid as Ltac2 +bindings, so that there will be a syntax error if one of the bound variables +starts with an uppercase character. + +The semantics of this construction is otherwise the same as the corresponding +one from Ltac1, except that it requires the goal to be focused. + +## Match over goals + +Similarly, there is a way to match over goals in an elegant way, which is +just a notation desugared at parsing time. + +``` +TERM ::= +| "match!" MATCH-ORDER? "goal" "with" GOAL-MATCHING* "end" +| "lazy_match!" MATCH-ORDER? "goal" "with" GOAL-MATCHING* "end" +| "multi_match!" MATCH-ORDER? "goal" "with" GOAL-MATCHING*"end" + +GOAL-MATCHING := +| "|" "[" HYP-MATCHING* "|-" CONSTR-PATTERN "]" "=>" TERM + +HYP-MATCHING := +| LIDENT ":" CONSTR-PATTERN + +MATCH-ORDER := +| "reverse" +``` + +Variables from `HYP-MATCHING` and `CONSTR-PATTERN` are bound in the body of the +branch. Their types are: +- `constr` for pattern variables appearing in a `CONSTR` +- `Pattern.context` for variables binding a context +- `ident` for variables binding a hypothesis name. + +The same identifier caveat as in the case of matching over constr applies, and +this features has the same semantics as in Ltac1. In particular, a `reverse` +flag can be specified to match hypotheses from the more recently introduced to +the least recently introduced one. + +# Notations + +Notations are the crux of the usability of Ltac1. We should be able to recover +a feeling similar to the old implementation by using and abusing notations. + +## Scopes + +A scope is a name given to a grammar entry used to produce some Ltac2 expression +at parsing time. Scopes are described using a form of S-expression. + +``` +SCOPE := +| STRING +| INT +| LIDENT ( "(" SCOPE₀ "," ... "," SCOPEₙ ")" ) +``` + +A few scopes contain antiquotation features. For sake of uniformity, all +antiquotations are introduced by the syntax `"$" VAR`. + +The following scopes are built-in. +- constr: + + parses `c = COQCONSTR` and produces `constr:(c)` +- ident: + + parses `id = IDENT` and produces `ident:(id)` + + parses `"$" (x = IDENT)` and produces the variable `x` +- list0(*scope*): + + if *scope* parses `ENTRY`, parses ̀`(x₀, ..., xₙ = ENTRY*)` and produces + `[x₀; ...; xₙ]`. +- list0(*scope*, sep = STRING): + + if *scope* parses `ENTRY`, parses `(x₀ = ENTRY, "sep", ..., "sep", xₙ = ENTRY)` + and produces `[x₀; ...; xₙ]`. +- list1: same as list0 (with or without separator) but parses `ENTRY+` instead + of `ENTRY*`. +- opt(*scope*) + + if *scope* parses `ENTRY`, parses `ENTRY?` and produces either `None` or + `Some x` where `x` is the parsed expression. +- self: + + parses a Ltac2 expression at the current level and return it as is. +- next: + + parses a Ltac2 expression at the next level and return it as is. +- tactic(n = INT): + + parses a Ltac2 expression at the provided level *n* and return it as is. +- thunk(*scope*): + + parses the same as *scope*, and if *e* is the parsed expression, returns + `fun () => e`. +- STRING: + + parses the corresponding string as a CAMLP5 IDENT and returns `()`. +- keyword(s = STRING): + + parses the string *s* as a keyword and returns `()`. +- terminal(s = STRING): + + parses the string *s* as a keyword, if it is already a + keyword, otherwise as an IDENT. Returns `()`. +- seq(*scope₁*, ..., *scopeₙ*): + + parses *scope₁*, ..., *scopeₙ* in this order, and produces a tuple made + out of the parsed values in the same order. As an optimization, all + subscopes of the form STRING are left out of the returned tuple, instead + of returning a useless unit value. It is forbidden for the various + subscopes to refer to the global entry using self of next. + +A few other specific scopes exist to handle Ltac1-like syntax, but their use is +discouraged and they are thus not documented. + +For now there is no way to declare new scopes from Ltac2 side, but this is +planned. + +## Notations + +The Ltac2 parser can be extended by syntactic notations. +``` +VERNAC ::= +| "Ltac2" "Notation" TOKEN+ LEVEL? ":=" TERM + +LEVEL := ":" INT + +TOKEN := +| VAR "(" SCOPE ")" +| STRING +``` + +A Ltac2 notation adds a parsing rule to the Ltac2 grammar, which is expanded +to the provided body where every token from the notation is let-bound to the +corresponding generated expression. + +For instance, assume we perform: +``` +Ltac2 Notation "foo" c(thunk(constr)) ids(list0(ident)) := Bar.f c ids. +``` +Then the following expression +``` +let y := @X in foo (nat -> nat) x $y +``` +will expand at parsing time to +``` +let y := @X in +let c := fun () => constr:(nat -> nat) with ids := [@x; y] in Bar.f c ids +``` + +Beware that the order of evaluation of multiple let-bindings is not specified, +so that you may have to resort to thunking to ensure that side-effects are +performed at the right time. + +## Abbreviations + +There exists a special kind of notations, called abbreviations, that is designed +so that it does not add any parsing rules. It is similar in spirit to Coq +abbreviations, insofar as its main purpose is to give an absolute name to a +piece of pure syntax, which can be transparently referred by this name as if it +were a proper definition. Abbreviations are introduced by the following +syntax. + +``` +VERNAC ::= +| "Ltac2" "Notation" IDENT ":=" TERM +``` + +The abbreviation can then be manipulated just as a normal Ltac2 definition, +except that it is expanded at internalization time into the given expression. +Furthermore, in order to make this kind of construction useful in practice in +an effectful language such as Ltac2, any syntactic argument to an abbreviation +is thunked on-the-fly during its expansion. + +For instance, suppose that we define the following. +``` +Ltac2 Notation foo := fun x => x (). +``` +Then we have the following expansion at internalization time. +``` +foo 0 ↦ (fun x => x ()) (fun _ => 0) +``` + +Note that abbreviations are not typechecked at all, and may result in typing +errors after expansion. + +# Evaluation + +Ltac2 features a toplevel loop that can be used to evaluate expressions. + +``` +VERNAC ::= +| "Ltac2" "Eval" TERM +``` + +This command evaluates the term in the current proof if there is one, or in the +global environment otherwise, and displays the resulting value to the user +together with its type. This function is pure in the sense that it does not +modify the state of the proof, and in particular all side-effects are discarded. + +# Debug + +When the option `Ltac2 Backtrace` is set, toplevel failures will be printed with +a backtrace. + +# Compatibility layer with Ltac1 + +## Ltac1 from Ltac2 + +### Simple API + +One can call Ltac1 code from Ltac2 by using the `ltac1` quotation. It parses +a Ltac1 expression, and semantics of this quotation is the evaluation of the +corresponding code for its side effects. In particular, in cannot return values, +and the quotation has type `unit`. + +Beware, Ltac1 **cannot** access variables from the Ltac2 scope. One is limited +to the use of standalone function calls. + +### Low-level API + +There exists a lower-level FFI into Ltac1 that is not recommended for daily use, +which is available in the `Ltac2.Ltac1` module. This API allows to directly +manipulate dynamically-typed Ltac1 values, either through the function calls, +or using the `ltac1val` quotation. The latter parses the same as `ltac1`, but +has type `Ltac2.Ltac1.t` instead of `unit`, and dynamically behaves as an Ltac1 +thunk, i.e. `ltac1val:(foo)` corresponds to the tactic closure that Ltac1 +would generate from `idtac; foo`. + +Due to intricate dynamic semantics, understanding when Ltac1 value quotations +focus is very hard. This is why some functions return a continuation-passing +style value, as it can dispatch dynamically between focused and unfocused +behaviour. + +## Ltac2 from Ltac1 + +Same as above by switching Ltac1 by Ltac2 and using the `ltac2` quotation +instead. + +Note that the tactic expression is evaluated eagerly, if one wants to use it as +an argument to a Ltac1 function, she has to resort to the good old +`idtac; ltac2:(foo)` trick. For instance, the code below will fail immediately +and won't print anything. + +``` +Ltac mytac tac := idtac "wow"; tac. + +Goal True. +Proof. +mytac ltac2:(fail). +``` + +# Transition from Ltac1 + +Owing to the use of a bunch of notations, the transition shouldn't be +atrociously horrible and shockingly painful up to the point you want to retire +in the Ariège mountains, living off the land and insulting careless bypassers in +proto-georgian. + +That said, we do *not* guarantee you it is going to be a blissful walk either. +Hopefully, owing to the fact Ltac2 is typed, the interactive dialogue with Coq +will help you. + +We list the major changes and the transition strategies hereafter. + +## Syntax changes + +Due to conflicts, a few syntactic rules have changed. + +- The dispatch tactical `tac; [foo|bar]` is now written `tac > [foo|bar]`. +- Levels of a few operators have been revised. Some tacticals now parse as if + they were a normal function, i.e. one has to put parentheses around the + argument when it is complex, e.g an abstraction. List of affected tacticals: + `try`, `repeat`, `do`, `once`, `progress`, `time`, `abstract`. +- `idtac` is no more. Either use `()` if you expect nothing to happen, + `(fun () => ())` if you want a thunk (see next section), or use printing + primitives from the `Message` module if you want to display something. + +## Tactic delay + +Tactics are not magically delayed anymore, neither as functions nor as +arguments. It is your responsibility to thunk them beforehand and apply them +at the call site. + +A typical example of a delayed function: +``` +Ltac foo := blah. +``` +becomes +``` +Ltac2 foo () := blah. +``` + +All subsequent calls to `foo` must be applied to perform the same effect as +before. + +Likewise, for arguments: +``` +Ltac bar tac := tac; tac; tac. +``` +becomes +``` +Ltac2 bar tac := tac (); tac (); tac (). +``` + +We recommend the use of syntactic notations to ease the transition. For +instance, the first example can alternatively written as: +``` +Ltac2 foo0 () := blah. +Ltac2 Notation foo := foo0 (). +``` + +This allows to keep the subsequent calls to the tactic as-is, as the +expression `foo` will be implicitly expanded everywhere into `foo0 ()`. Such +a trick also works for arguments, as arguments of syntactic notations are +implicitly thunked. The second example could thus be written as follows. + +``` +Ltac2 bar0 tac := tac (); tac (); tac (). +Ltac2 Notation bar := bar0. +``` + +## Variable binding + +Ltac1 relies on a crazy amount of dynamic trickery to be able to tell apart +bound variables from terms, hypotheses and whatnot. There is no such thing in +Ltac2, as variables are recognized statically and other constructions do not +live in the same syntactic world. Due to the abuse of quotations, it can +sometimes be complicated to know what a mere identifier represents in a tactic +expression. We recommend tracking the context and letting the compiler spit +typing errors to understand what is going on. + +We list below the typical changes one has to perform depending on the static +errors produced by the typechecker. + +### In Ltac expressions + +- `Unbound value X`, `Unbound constructor X`: + * if `X` is meant to be a term from the current stactic environment, replace + the problematic use by `'X`. + * if `X` is meant to be a hypothesis from the goal context, replace the + problematic use by `&X`. + +### In quotations + +- `The reference X was not found in the current environment`: + * if `X` is meant to be a tactic expression bound by a Ltac2 let or function, + replace the problematic use by `$X`. + * if `X` is meant to be a hypothesis from the goal context, replace the + problematic use by `&X`. + +## Exception catching + +Ltac2 features a proper exception-catching mechanism. For this reason, the +Ltac1 mechanism relying on `fail` taking integers and tacticals decreasing it +has been removed. Now exceptions are preserved by all tacticals, and it is +your duty to catch it and reraise it depending on your use. + +# TODO + +- Implement deep pattern-matching. +- Craft an expressive set of primitive functions +- Implement native compilation to OCaml diff --git a/vendor/Ltac2/dune b/vendor/Ltac2/dune new file mode 100644 index 0000000000..5dbc4db66a --- /dev/null +++ b/vendor/Ltac2/dune @@ -0,0 +1,3 @@ +(env + (dev (flags :standard -rectypes)) + (release (flags :standard -rectypes))) diff --git a/vendor/Ltac2/dune-project b/vendor/Ltac2/dune-project new file mode 100644 index 0000000000..8154e999de --- /dev/null +++ b/vendor/Ltac2/dune-project @@ -0,0 +1,3 @@ +(lang dune 1.6) +(using coq 0.1) +(name ltac2) diff --git a/vendor/Ltac2/ltac2.opam b/vendor/Ltac2/ltac2.opam new file mode 100644 index 0000000000..47ceb882b1 --- /dev/null +++ b/vendor/Ltac2/ltac2.opam @@ -0,0 +1,18 @@ +synopsis: "A Tactic Language for Coq." +description: "A Tactic Language for Coq." +name: "coq-ltac2" +opam-version: "2.0" +maintainer: "Pierre-Marie Pédrot " +authors: "Pierre-Marie Pédrot " +homepage: "https://github.com/ppedrot/ltac2" +dev-repo: "https://github.com/ppedrot/ltac2.git" +bug-reports: "https://github.com/ppedrot/ltac2/issues" +license: "LGPL 2.1" +doc: "https://ppedrot.github.io/ltac2/doc" + +depends: [ + "coq" { = "dev" } + "dune" { build & >= "1.9.0" } +] + +build: [ "dune" "build" "-p" name "-j" jobs ] diff --git a/vendor/Ltac2/src/dune b/vendor/Ltac2/src/dune new file mode 100644 index 0000000000..332f3644b0 --- /dev/null +++ b/vendor/Ltac2/src/dune @@ -0,0 +1,11 @@ +(library + (name ltac2_plugin) + (public_name ltac2.plugin) + (modules_without_implementation tac2expr tac2qexpr tac2types) + (flags :standard -warn-error -9-27-50) + (libraries coq.plugins.firstorder)) + +(rule + (targets g_ltac2.ml) + (deps (:mlg-file g_ltac2.mlg)) + (action (run coqpp %{mlg-file}))) diff --git a/vendor/Ltac2/src/g_ltac2.mlg b/vendor/Ltac2/src/g_ltac2.mlg new file mode 100644 index 0000000000..0071dbb088 --- /dev/null +++ b/vendor/Ltac2/src/g_ltac2.mlg @@ -0,0 +1,933 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Tok.t Stream.t -> int option + +let entry_of_lookahead s (lk : lookahead) = + let run strm = match lk 0 strm with None -> err () | Some _ -> () in + Pcoq.Entry.of_parser s run + +let (>>) (lk1 : lookahead) lk2 n strm = match lk1 n strm with +| None -> None +| Some n -> lk2 n strm + +let (<+>) (lk1 : lookahead) lk2 n strm = match lk1 n strm with +| None -> lk2 n strm +| Some n -> Some n + +let lk_kw kw n strm = match stream_nth n strm with +| KEYWORD kw' | IDENT kw' -> if String.equal kw kw' then Some (n + 1) else None +| _ -> None + +let lk_ident n strm = match stream_nth n strm with +| IDENT _ -> Some (n + 1) +| _ -> None + +let lk_int n strm = match stream_nth n strm with +| NUMERAL { NumTok.int = _; frac = ""; exp = "" } -> Some (n + 1) +| _ -> None + +let lk_ident_or_anti = lk_ident <+> (lk_kw "$" >> lk_ident) + +(* lookahead for (x:=t), (?x:=t) and (1:=t) *) +let test_lpar_idnum_coloneq = + entry_of_lookahead "test_lpar_idnum_coloneq" begin + lk_kw "(" >> (lk_ident_or_anti <+> lk_int) >> lk_kw ":=" + end + +(* lookahead for (x:t), (?x:t) *) +let test_lpar_id_colon = + entry_of_lookahead "test_lpar_id_colon" begin + lk_kw "(" >> lk_ident_or_anti >> lk_kw ":" + end + +(* Hack to recognize "(x := t)" and "($x := t)" *) +let test_lpar_id_coloneq = + entry_of_lookahead "test_lpar_id_coloneq" begin + lk_kw "(" >> lk_ident_or_anti >> lk_kw ":=" + end + +(* Hack to recognize "(x)" *) +let test_lpar_id_rpar = + entry_of_lookahead "test_lpar_id_rpar" begin + lk_kw "(" >> lk_ident >> lk_kw ")" + end + +let test_ampersand_ident = + entry_of_lookahead "test_ampersand_ident" begin + lk_kw "&" >> lk_ident + end + +let test_dollar_ident = + entry_of_lookahead "test_dollar_ident" begin + lk_kw "$" >> lk_ident + end + +let tac2expr = Tac2entries.Pltac.tac2expr +let tac2type = Entry.create "tactic:tac2type" +let tac2def_val = Entry.create "tactic:tac2def_val" +let tac2def_typ = Entry.create "tactic:tac2def_typ" +let tac2def_ext = Entry.create "tactic:tac2def_ext" +let tac2def_syn = Entry.create "tactic:tac2def_syn" +let tac2def_mut = Entry.create "tactic:tac2def_mut" +let tac2def_run = Entry.create "tactic:tac2def_run" +let tac2mode = Entry.create "vernac:ltac2_command" + +let ltac1_expr = Pltac.tactic_expr + +let inj_wit wit loc x = CAst.make ~loc @@ CTacExt (wit, x) +let inj_open_constr loc c = inj_wit Tac2quote.wit_open_constr loc c +let inj_pattern loc c = inj_wit Tac2quote.wit_pattern loc c +let inj_reference loc c = inj_wit Tac2quote.wit_reference loc c +let inj_ltac1 loc e = inj_wit Tac2quote.wit_ltac1 loc e +let inj_ltac1val loc e = inj_wit Tac2quote.wit_ltac1val loc e + +let pattern_of_qualid qid = + if Tac2env.is_constructor qid then CAst.make ?loc:qid.CAst.loc @@ CPatRef (RelId qid, []) + else + let open Libnames in + if qualid_is_ident qid then CAst.make ?loc:qid.CAst.loc @@ CPatVar (Name (qualid_basename qid)) + else + CErrors.user_err ?loc:qid.CAst.loc (Pp.str "Syntax error") + +} + +GRAMMAR EXTEND Gram + GLOBAL: tac2expr tac2type tac2def_val tac2def_typ tac2def_ext tac2def_syn + tac2def_mut tac2def_run; + tac2pat: + [ "1" LEFTA + [ qid = Prim.qualid; pl = LIST1 tac2pat LEVEL "0" -> { + if Tac2env.is_constructor qid then + CAst.make ~loc @@ CPatRef (RelId qid, pl) + else + CErrors.user_err ~loc (Pp.str "Syntax error") } + | qid = Prim.qualid -> { pattern_of_qualid qid } + | "["; "]" -> { CAst.make ~loc @@ CPatRef (AbsKn (Other Tac2core.Core.c_nil), []) } + | p1 = tac2pat; "::"; p2 = tac2pat -> + { CAst.make ~loc @@ CPatRef (AbsKn (Other Tac2core.Core.c_cons), [p1; p2])} + ] + | "0" + [ "_" -> { CAst.make ~loc @@ CPatVar Anonymous } + | "()" -> { CAst.make ~loc @@ CPatRef (AbsKn (Tuple 0), []) } + | qid = Prim.qualid -> { pattern_of_qualid qid } + | "("; p = atomic_tac2pat; ")" -> { p } + ] ] + ; + atomic_tac2pat: + [ [ -> + { CAst.make ~loc @@ CPatRef (AbsKn (Tuple 0), []) } + | p = tac2pat; ":"; t = tac2type -> + { CAst.make ~loc @@ CPatCnv (p, t) } + | p = tac2pat; ","; pl = LIST0 tac2pat SEP "," -> + { let pl = p :: pl in + CAst.make ~loc @@ CPatRef (AbsKn (Tuple (List.length pl)), pl) } + | p = tac2pat -> { p } + ] ] + ; + tac2expr: + [ "6" RIGHTA + [ e1 = SELF; ";"; e2 = SELF -> { CAst.make ~loc @@ CTacSeq (e1, e2) } ] + | "5" + [ "fun"; it = LIST1 input_fun ; "=>"; body = tac2expr LEVEL "6" -> + { CAst.make ~loc @@ CTacFun (it, body) } + | "let"; isrec = rec_flag; + lc = LIST1 let_clause SEP "with"; "in"; + e = tac2expr LEVEL "6" -> + { CAst.make ~loc @@ CTacLet (isrec, lc, e) } + | "match"; e = tac2expr LEVEL "5"; "with"; bl = branches; "end" -> + { CAst.make ~loc @@ CTacCse (e, bl) } + ] + | "4" LEFTA [ ] + | "::" RIGHTA + [ e1 = tac2expr; "::"; e2 = tac2expr -> + { CAst.make ~loc @@ CTacApp (CAst.make ~loc @@ CTacCst (AbsKn (Other Tac2core.Core.c_cons)), [e1; e2]) } + ] + | [ e0 = SELF; ","; el = LIST1 NEXT SEP "," -> + { let el = e0 :: el in + CAst.make ~loc @@ CTacApp (CAst.make ~loc @@ CTacCst (AbsKn (Tuple (List.length el))), el) } ] + | "1" LEFTA + [ e = tac2expr; el = LIST1 tac2expr LEVEL "0" -> + { CAst.make ~loc @@ CTacApp (e, el) } + | e = SELF; ".("; qid = Prim.qualid; ")" -> + { CAst.make ~loc @@ CTacPrj (e, RelId qid) } + | e = SELF; ".("; qid = Prim.qualid; ")"; ":="; r = tac2expr LEVEL "5" -> + { CAst.make ~loc @@ CTacSet (e, RelId qid, r) } ] + | "0" + [ "("; a = SELF; ")" -> { a } + | "("; a = SELF; ":"; t = tac2type; ")" -> + { CAst.make ~loc @@ CTacCnv (a, t) } + | "()" -> + { CAst.make ~loc @@ CTacCst (AbsKn (Tuple 0)) } + | "("; ")" -> + { CAst.make ~loc @@ CTacCst (AbsKn (Tuple 0)) } + | "["; a = LIST0 tac2expr LEVEL "5" SEP ";"; "]" -> + { Tac2quote.of_list ~loc (fun x -> x) a } + | "{"; a = tac2rec_fieldexprs; "}" -> + { CAst.make ~loc @@ CTacRec a } + | a = tactic_atom -> { a } ] + ] + ; + branches: + [ [ -> { [] } + | "|"; bl = LIST1 branch SEP "|" -> { bl } + | bl = LIST1 branch SEP "|" -> { bl } ] + ] + ; + branch: + [ [ pat = tac2pat LEVEL "1"; "=>"; e = tac2expr LEVEL "6" -> { (pat, e) } ] ] + ; + rec_flag: + [ [ IDENT "rec" -> { true } + | -> { false } ] ] + ; + mut_flag: + [ [ IDENT "mutable" -> { true } + | -> { false } ] ] + ; + typ_param: + [ [ "'"; id = Prim.ident -> { id } ] ] + ; + tactic_atom: + [ [ n = Prim.integer -> { CAst.make ~loc @@ CTacAtm (AtmInt n) } + | s = Prim.string -> { CAst.make ~loc @@ CTacAtm (AtmStr s) } + | qid = Prim.qualid -> + { if Tac2env.is_constructor qid then + CAst.make ~loc @@ CTacCst (RelId qid) + else + CAst.make ~loc @@ CTacRef (RelId qid) } + | "@"; id = Prim.ident -> { Tac2quote.of_ident (CAst.make ~loc id) } + | "&"; id = lident -> { Tac2quote.of_hyp ~loc id } + | "'"; c = Constr.constr -> { inj_open_constr loc c } + | IDENT "constr"; ":"; "("; c = Constr.lconstr; ")" -> { Tac2quote.of_constr c } + | IDENT "open_constr"; ":"; "("; c = Constr.lconstr; ")" -> { Tac2quote.of_open_constr c } + | IDENT "ident"; ":"; "("; c = lident; ")" -> { Tac2quote.of_ident c } + | IDENT "pattern"; ":"; "("; c = Constr.lconstr_pattern; ")" -> { inj_pattern loc c } + | IDENT "reference"; ":"; "("; c = globref; ")" -> { inj_reference loc c } + | IDENT "ltac1"; ":"; "("; qid = ltac1_expr; ")" -> { inj_ltac1 loc qid } + | IDENT "ltac1val"; ":"; "("; qid = ltac1_expr; ")" -> { inj_ltac1val loc qid } + ] ] + ; + let_clause: + [ [ binder = let_binder; ":="; te = tac2expr -> + { let (pat, fn) = binder in + let te = match fn with + | None -> te + | Some args -> CAst.make ~loc @@ CTacFun (args, te) + in + (pat, te) } + ] ] + ; + let_binder: + [ [ pats = LIST1 input_fun -> + { match pats with + | [{CAst.v=CPatVar _} as pat] -> (pat, None) + | ({CAst.v=CPatVar (Name id)} as pat) :: args -> (pat, Some args) + | [pat] -> (pat, None) + | _ -> CErrors.user_err ~loc (str "Invalid pattern") } + ] ] + ; + tac2type: + [ "5" RIGHTA + [ t1 = tac2type; "->"; t2 = tac2type -> { CAst.make ~loc @@ CTypArrow (t1, t2) } ] + | "2" + [ t = tac2type; "*"; tl = LIST1 tac2type LEVEL "1" SEP "*" -> + { let tl = t :: tl in + CAst.make ~loc @@ CTypRef (AbsKn (Tuple (List.length tl)), tl) } ] + | "1" LEFTA + [ t = SELF; qid = Prim.qualid -> { CAst.make ~loc @@ CTypRef (RelId qid, [t]) } ] + | "0" + [ "("; t = tac2type LEVEL "5"; ")" -> { t } + | id = typ_param -> { CAst.make ~loc @@ CTypVar (Name id) } + | "_" -> { CAst.make ~loc @@ CTypVar Anonymous } + | qid = Prim.qualid -> { CAst.make ~loc @@ CTypRef (RelId qid, []) } + | "("; p = LIST1 tac2type LEVEL "5" SEP ","; ")"; qid = Prim.qualid -> + { CAst.make ~loc @@ CTypRef (RelId qid, p) } ] + ]; + locident: + [ [ id = Prim.ident -> { CAst.make ~loc id } ] ] + ; + binder: + [ [ "_" -> { CAst.make ~loc Anonymous } + | l = Prim.ident -> { CAst.make ~loc (Name l) } ] ] + ; + input_fun: + [ [ b = tac2pat LEVEL "0" -> { b } ] ] + ; + tac2def_body: + [ [ name = binder; it = LIST0 input_fun; ":="; e = tac2expr -> + { let e = if List.is_empty it then e else CAst.make ~loc @@ CTacFun (it, e) in + (name, e) } + ] ] + ; + tac2def_val: + [ [ mut = mut_flag; isrec = rec_flag; l = LIST1 tac2def_body SEP "with" -> + { StrVal (mut, isrec, l) } + ] ] + ; + tac2def_mut: + [ [ "Set"; qid = Prim.qualid; ":="; e = tac2expr -> { StrMut (qid, e) } ] ] + ; + tac2def_run: + [ [ "Eval"; e = tac2expr -> { StrRun e } ] ] + ; + tac2typ_knd: + [ [ t = tac2type -> { CTydDef (Some t) } + | "["; ".."; "]" -> { CTydOpn } + | "["; t = tac2alg_constructors; "]" -> { CTydAlg t } + | "{"; t = tac2rec_fields; "}"-> { CTydRec t } ] ] + ; + tac2alg_constructors: + [ [ "|"; cs = LIST1 tac2alg_constructor SEP "|" -> { cs } + | cs = LIST0 tac2alg_constructor SEP "|" -> { cs } ] ] + ; + tac2alg_constructor: + [ [ c = Prim.ident -> { (c, []) } + | c = Prim.ident; "("; args = LIST0 tac2type SEP ","; ")"-> { (c, args) } ] ] + ; + tac2rec_fields: + [ [ f = tac2rec_field; ";"; l = tac2rec_fields -> { f :: l } + | f = tac2rec_field; ";" -> { [f] } + | f = tac2rec_field -> { [f] } + | -> { [] } ] ] + ; + tac2rec_field: + [ [ mut = mut_flag; id = Prim.ident; ":"; t = tac2type -> { (id, mut, t) } ] ] + ; + tac2rec_fieldexprs: + [ [ f = tac2rec_fieldexpr; ";"; l = tac2rec_fieldexprs -> { f :: l } + | f = tac2rec_fieldexpr; ";" -> { [f] } + | f = tac2rec_fieldexpr-> { [f] } + | -> { [] } ] ] + ; + tac2rec_fieldexpr: + [ [ qid = Prim.qualid; ":="; e = tac2expr LEVEL "1" -> { RelId qid, e } ] ] + ; + tac2typ_prm: + [ [ -> { [] } + | id = typ_param -> { [CAst.make ~loc id] } + | "("; ids = LIST1 [ id = typ_param -> { CAst.make ~loc id } ] SEP "," ;")" -> { ids } + ] ] + ; + tac2typ_def: + [ [ prm = tac2typ_prm; id = Prim.qualid; b = tac2type_body -> { let (r, e) = b in (id, r, (prm, e)) } ] ] + ; + tac2type_body: + [ [ -> { false, CTydDef None } + | ":="; e = tac2typ_knd -> { false, e } + | "::="; e = tac2typ_knd -> { true, e } + ] ] + ; + tac2def_typ: + [ [ "Type"; isrec = rec_flag; l = LIST1 tac2typ_def SEP "with" -> + { StrTyp (isrec, l) } + ] ] + ; + tac2def_ext: + [ [ "@"; IDENT "external"; id = locident; ":"; t = tac2type LEVEL "5"; ":="; + plugin = Prim.string; name = Prim.string -> + { let ml = { mltac_plugin = plugin; mltac_tactic = name } in + StrPrm (id, t, ml) } + ] ] + ; + syn_node: + [ [ "_" -> { CAst.make ~loc None } + | id = Prim.ident -> { CAst.make ~loc (Some id) } + ] ] + ; + sexpr: + [ [ s = Prim.string -> { SexprStr (CAst.make ~loc s) } + | n = Prim.integer -> { SexprInt (CAst.make ~loc n) } + | id = syn_node -> { SexprRec (loc, id, []) } + | id = syn_node; "("; tok = LIST1 sexpr SEP "," ; ")" -> + { SexprRec (loc, id, tok) } + ] ] + ; + syn_level: + [ [ -> { None } + | ":"; n = Prim.integer -> { Some n } + ] ] + ; + tac2def_syn: + [ [ "Notation"; toks = LIST1 sexpr; n = syn_level; ":="; + e = tac2expr -> + { StrSyn (toks, n, e) } + ] ] + ; + lident: + [ [ id = Prim.ident -> { CAst.make ~loc id } ] ] + ; + globref: + [ [ "&"; id = Prim.ident -> { CAst.make ~loc (QHypothesis id) } + | qid = Prim.qualid -> { CAst.make ~loc @@ QReference qid } + ] ] + ; +END + +(* Quotation scopes used by notations *) + +{ + +open Tac2entries.Pltac + +let loc_of_ne_list l = Loc.merge_opt (fst (List.hd l)) (fst (List.last l)) + +} + +GRAMMAR EXTEND Gram + GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause + q_conversion q_rewriting q_clause q_dispatch q_occurrences q_strategy_flag + q_destruction_arg q_reference q_with_bindings q_constr_matching + q_goal_matching q_hintdb q_move_location q_pose q_assert; + anti: + [ [ "$"; id = Prim.ident -> { QAnti (CAst.make ~loc id) } ] ] + ; + ident_or_anti: + [ [ id = lident -> { QExpr id } + | "$"; id = Prim.ident -> { QAnti (CAst.make ~loc id) } + ] ] + ; + lident: + [ [ id = Prim.ident -> { CAst.make ~loc id } ] ] + ; + lnatural: + [ [ n = Prim.natural -> { CAst.make ~loc n } ] ] + ; + q_ident: + [ [ id = ident_or_anti -> { id } ] ] + ; + qhyp: + [ [ x = anti -> { x } + | n = lnatural -> { QExpr (CAst.make ~loc @@ QAnonHyp n) } + | id = lident -> { QExpr (CAst.make ~loc @@ QNamedHyp id) } + ] ] + ; + simple_binding: + [ [ "("; h = qhyp; ":="; c = Constr.lconstr; ")" -> + { CAst.make ~loc (h, c) } + ] ] + ; + bindings: + [ [ test_lpar_idnum_coloneq; bl = LIST1 simple_binding -> + { CAst.make ~loc @@ QExplicitBindings bl } + | bl = LIST1 Constr.constr -> + { CAst.make ~loc @@ QImplicitBindings bl } + ] ] + ; + q_bindings: + [ [ bl = bindings -> { bl } ] ] + ; + q_with_bindings: + [ [ bl = with_bindings -> { bl } ] ] + ; + intropatterns: + [ [ l = LIST0 nonsimple_intropattern -> { CAst.make ~loc l } ] ] + ; +(* ne_intropatterns: *) +(* [ [ l = LIST1 nonsimple_intropattern -> l ]] *) +(* ; *) + or_and_intropattern: + [ [ "["; tc = LIST1 intropatterns SEP "|"; "]" -> { CAst.make ~loc @@ QIntroOrPattern tc } + | "()" -> { CAst.make ~loc @@ QIntroAndPattern (CAst.make ~loc []) } + | "("; si = simple_intropattern; ")" -> { CAst.make ~loc @@ QIntroAndPattern (CAst.make ~loc [si]) } + | "("; si = simple_intropattern; ","; + tc = LIST1 simple_intropattern SEP "," ; ")" -> + { CAst.make ~loc @@ QIntroAndPattern (CAst.make ~loc (si::tc)) } + | "("; si = simple_intropattern; "&"; + tc = LIST1 simple_intropattern SEP "&" ; ")" -> + (* (A & B & C) is translated into (A,(B,C)) *) + { let rec pairify = function + | ([]|[_]|[_;_]) as l -> CAst.make ~loc l + | t::q -> + let q = + CAst.make ~loc @@ + QIntroAction (CAst.make ~loc @@ + QIntroOrAndPattern (CAst.make ~loc @@ + QIntroAndPattern (pairify q))) + in + CAst.make ~loc [t; q] + in CAst.make ~loc @@ QIntroAndPattern (pairify (si::tc)) } ] ] + ; + equality_intropattern: + [ [ "->" -> { CAst.make ~loc @@ QIntroRewrite true } + | "<-" -> { CAst.make ~loc @@ QIntroRewrite false } + | "[="; tc = intropatterns; "]" -> { CAst.make ~loc @@ QIntroInjection tc } ] ] + ; + naming_intropattern: + [ [ LEFTQMARK; id = lident -> + { CAst.make ~loc @@ QIntroFresh (QExpr id) } + | "?$"; id = lident -> + { CAst.make ~loc @@ QIntroFresh (QAnti id) } + | "?" -> + { CAst.make ~loc @@ QIntroAnonymous } + | id = ident_or_anti -> + { CAst.make ~loc @@ QIntroIdentifier id } + ] ] + ; + nonsimple_intropattern: + [ [ l = simple_intropattern -> { l } + | "*" -> { CAst.make ~loc @@ QIntroForthcoming true } + | "**" -> { CAst.make ~loc @@ QIntroForthcoming false } ] ] + ; + simple_intropattern: + [ [ pat = simple_intropattern_closed -> +(* l = LIST0 ["%"; c = operconstr LEVEL "0" -> c] -> *) + (** TODO: handle %pat *) + { pat } + ] ] + ; + simple_intropattern_closed: + [ [ pat = or_and_intropattern -> + { CAst.make ~loc @@ QIntroAction (CAst.make ~loc @@ QIntroOrAndPattern pat) } + | pat = equality_intropattern -> + { CAst.make ~loc @@ QIntroAction pat } + | "_" -> + { CAst.make ~loc @@ QIntroAction (CAst.make ~loc @@ QIntroWildcard) } + | pat = naming_intropattern -> + { CAst.make ~loc @@ QIntroNaming pat } + ] ] + ; + q_intropatterns: + [ [ ipat = intropatterns -> { ipat } ] ] + ; + q_intropattern: + [ [ ipat = simple_intropattern -> { ipat } ] ] + ; + nat_or_anti: + [ [ n = lnatural -> { QExpr n } + | "$"; id = Prim.ident -> { QAnti (CAst.make ~loc id) } + ] ] + ; + eqn_ipat: + [ [ IDENT "eqn"; ":"; pat = naming_intropattern -> { Some pat } + | -> { None } + ] ] + ; + with_bindings: + [ [ "with"; bl = bindings -> { bl } | -> { CAst.make ~loc @@ QNoBindings } ] ] + ; + constr_with_bindings: + [ [ c = Constr.constr; l = with_bindings -> { CAst.make ~loc @@ (c, l) } ] ] + ; + destruction_arg: + [ [ n = lnatural -> { CAst.make ~loc @@ QElimOnAnonHyp n } + | id = lident -> { CAst.make ~loc @@ QElimOnIdent id } + | c = constr_with_bindings -> { CAst.make ~loc @@ QElimOnConstr c } + ] ] + ; + q_destruction_arg: + [ [ arg = destruction_arg -> { arg } ] ] + ; + as_or_and_ipat: + [ [ "as"; ipat = or_and_intropattern -> { Some ipat } + | -> { None } + ] ] + ; + occs_nums: + [ [ nl = LIST1 nat_or_anti -> { CAst.make ~loc @@ QOnlyOccurrences nl } + | "-"; n = nat_or_anti; nl = LIST0 nat_or_anti -> + { CAst.make ~loc @@ QAllOccurrencesBut (n::nl) } + ] ] + ; + occs: + [ [ "at"; occs = occs_nums -> { occs } | -> { CAst.make ~loc QAllOccurrences } ] ] + ; + hypident: + [ [ id = ident_or_anti -> + { id,Locus.InHyp } + | "("; IDENT "type"; IDENT "of"; id = ident_or_anti; ")" -> + { id,Locus.InHypTypeOnly } + | "("; IDENT "value"; IDENT "of"; id = ident_or_anti; ")" -> + { id,Locus.InHypValueOnly } + ] ] + ; + hypident_occ: + [ [ h=hypident; occs=occs -> { let (id,l) = h in ((occs,id),l) } ] ] + ; + in_clause: + [ [ "*"; occs=occs -> + { { q_onhyps = None; q_concl_occs = occs } } + | "*"; "|-"; occs = concl_occ -> + { { q_onhyps = None; q_concl_occs = occs } } + | hl = LIST0 hypident_occ SEP ","; "|-"; occs = concl_occ -> + { { q_onhyps = Some hl; q_concl_occs = occs } } + | hl = LIST0 hypident_occ SEP "," -> + { { q_onhyps = Some hl; q_concl_occs = CAst.make ~loc QNoOccurrences } } + ] ] + ; + clause: + [ [ "in"; cl = in_clause -> { CAst.make ~loc @@ cl } + | "at"; occs = occs_nums -> + { CAst.make ~loc @@ { q_onhyps = Some []; q_concl_occs = occs } } + ] ] + ; + q_clause: + [ [ cl = clause -> { cl } ] ] + ; + concl_occ: + [ [ "*"; occs = occs -> { occs } + | -> { CAst.make ~loc QNoOccurrences } + ] ] + ; + induction_clause: + [ [ c = destruction_arg; pat = as_or_and_ipat; eq = eqn_ipat; + cl = OPT clause -> + { CAst.make ~loc @@ { + indcl_arg = c; + indcl_eqn = eq; + indcl_as = pat; + indcl_in = cl; + } } + ] ] + ; + q_induction_clause: + [ [ cl = induction_clause -> { cl } ] ] + ; + conversion: + [ [ c = Constr.constr -> + { CAst.make ~loc @@ QConvert c } + | c1 = Constr.constr; "with"; c2 = Constr.constr -> + { CAst.make ~loc @@ QConvertWith (c1, c2) } + ] ] + ; + q_conversion: + [ [ c = conversion -> { c } ] ] + ; + orient: + [ [ "->" -> { CAst.make ~loc (Some true) } + | "<-" -> { CAst.make ~loc (Some false) } + | -> { CAst.make ~loc None } + ]] + ; + rewriter: + [ [ "!"; c = constr_with_bindings -> + { (CAst.make ~loc @@ QRepeatPlus,c) } + | [ "?" -> { () } | LEFTQMARK -> { () } ]; c = constr_with_bindings -> + { (CAst.make ~loc @@ QRepeatStar,c) } + | n = lnatural; "!"; c = constr_with_bindings -> + { (CAst.make ~loc @@ QPrecisely n,c) } + | n = lnatural; ["?" -> { () } | LEFTQMARK -> { () } ]; c = constr_with_bindings -> + { (CAst.make ~loc @@ QUpTo n,c) } + | n = lnatural; c = constr_with_bindings -> + { (CAst.make ~loc @@ QPrecisely n,c) } + | c = constr_with_bindings -> + { (CAst.make ~loc @@ QPrecisely (CAst.make 1), c) } + ] ] + ; + oriented_rewriter: + [ [ b = orient; r = rewriter -> + { let (m, c) = r in + CAst.make ~loc @@ { + rew_orient = b; + rew_repeat = m; + rew_equatn = c; + } } + ] ] + ; + q_rewriting: + [ [ r = oriented_rewriter -> { r } ] ] + ; + tactic_then_last: + [ [ "|"; lta = LIST0 (OPT tac2expr LEVEL "6") SEP "|" -> { lta } + | -> { [] } + ] ] + ; + tactic_then_gen: + [ [ ta = tac2expr; "|"; tg = tactic_then_gen -> { let (first,last) = tg in (Some ta :: first, last) } + | ta = tac2expr; ".."; l = tactic_then_last -> { ([], Some (Some ta, l)) } + | ".."; l = tactic_then_last -> { ([], Some (None, l)) } + | ta = tac2expr -> { ([Some ta], None) } + | "|"; tg = tactic_then_gen -> { let (first,last) = tg in (None :: first, last) } + | -> { ([None], None) } + ] ] + ; + q_dispatch: + [ [ d = tactic_then_gen -> { CAst.make ~loc d } ] ] + ; + q_occurrences: + [ [ occs = occs -> { occs } ] ] + ; + red_flag: + [ [ IDENT "beta" -> { CAst.make ~loc @@ QBeta } + | IDENT "iota" -> { CAst.make ~loc @@ QIota } + | IDENT "match" -> { CAst.make ~loc @@ QMatch } + | IDENT "fix" -> { CAst.make ~loc @@ QFix } + | IDENT "cofix" -> { CAst.make ~loc @@ QCofix } + | IDENT "zeta" -> { CAst.make ~loc @@ QZeta } + | IDENT "delta"; d = delta_flag -> { d } + ] ] + ; + refglobal: + [ [ "&"; id = Prim.ident -> { QExpr (CAst.make ~loc @@ QHypothesis id) } + | qid = Prim.qualid -> { QExpr (CAst.make ~loc @@ QReference qid) } + | "$"; id = Prim.ident -> { QAnti (CAst.make ~loc id) } + ] ] + ; + q_reference: + [ [ r = refglobal -> { r } ] ] + ; + refglobals: + [ [ gl = LIST1 refglobal -> { CAst.make ~loc gl } ] ] + ; + delta_flag: + [ [ "-"; "["; idl = refglobals; "]" -> { CAst.make ~loc @@ QDeltaBut idl } + | "["; idl = refglobals; "]" -> { CAst.make ~loc @@ QConst idl } + | -> { CAst.make ~loc @@ QDeltaBut (CAst.make ~loc []) } + ] ] + ; + strategy_flag: + [ [ s = LIST1 red_flag -> { CAst.make ~loc s } + | d = delta_flag -> + { CAst.make ~loc + [CAst.make ~loc QBeta; CAst.make ~loc QIota; CAst.make ~loc QZeta; d] } + ] ] + ; + q_strategy_flag: + [ [ flag = strategy_flag -> { flag } ] ] + ; + hintdb: + [ [ "*" -> { CAst.make ~loc @@ QHintAll } + | l = LIST1 ident_or_anti -> { CAst.make ~loc @@ QHintDbs l } + ] ] + ; + q_hintdb: + [ [ db = hintdb -> { db } ] ] + ; + match_pattern: + [ [ IDENT "context"; id = OPT Prim.ident; + "["; pat = Constr.lconstr_pattern; "]" -> { CAst.make ~loc @@ QConstrMatchContext (id, pat) } + | pat = Constr.lconstr_pattern -> { CAst.make ~loc @@ QConstrMatchPattern pat } ] ] + ; + match_rule: + [ [ mp = match_pattern; "=>"; tac = tac2expr -> + { CAst.make ~loc @@ (mp, tac) } + ] ] + ; + match_list: + [ [ mrl = LIST1 match_rule SEP "|" -> { CAst.make ~loc @@ mrl } + | "|"; mrl = LIST1 match_rule SEP "|" -> { CAst.make ~loc @@ mrl } ] ] + ; + q_constr_matching: + [ [ m = match_list -> { m } ] ] + ; + gmatch_hyp_pattern: + [ [ na = Prim.name; ":"; pat = match_pattern -> { (na, pat) } ] ] + ; + gmatch_pattern: + [ [ "["; hl = LIST0 gmatch_hyp_pattern SEP ","; "|-"; p = match_pattern; "]" -> + { CAst.make ~loc @@ { + q_goal_match_concl = p; + q_goal_match_hyps = hl; + } } + ] ] + ; + gmatch_rule: + [ [ mp = gmatch_pattern; "=>"; tac = tac2expr -> + { CAst.make ~loc @@ (mp, tac) } + ] ] + ; + gmatch_list: + [ [ mrl = LIST1 gmatch_rule SEP "|" -> { CAst.make ~loc @@ mrl } + | "|"; mrl = LIST1 gmatch_rule SEP "|" -> { CAst.make ~loc @@ mrl } ] ] + ; + q_goal_matching: + [ [ m = gmatch_list -> { m } ] ] + ; + move_location: + [ [ "at"; IDENT "top" -> { CAst.make ~loc @@ QMoveFirst } + | "at"; IDENT "bottom" -> { CAst.make ~loc @@ QMoveLast } + | IDENT "after"; id = ident_or_anti -> { CAst.make ~loc @@ QMoveAfter id } + | IDENT "before"; id = ident_or_anti -> { CAst.make ~loc @@ QMoveBefore id } + ] ] + ; + q_move_location: + [ [ mv = move_location -> { mv } ] ] + ; + as_name: + [ [ -> { None } + | "as"; id = ident_or_anti -> { Some id } + ] ] + ; + pose: + [ [ test_lpar_id_coloneq; "("; id = ident_or_anti; ":="; c = Constr.lconstr; ")" -> + { CAst.make ~loc (Some id, c) } + | c = Constr.constr; na = as_name -> { CAst.make ~loc (na, c) } + ] ] + ; + q_pose: + [ [ p = pose -> { p } ] ] + ; + as_ipat: + [ [ "as"; ipat = simple_intropattern -> { Some ipat } + | -> { None } + ] ] + ; + by_tactic: + [ [ "by"; tac = tac2expr -> { Some tac } + | -> { None } + ] ] + ; + assertion: + [ [ test_lpar_id_coloneq; "("; id = ident_or_anti; ":="; c = Constr.lconstr; ")" -> + { CAst.make ~loc (QAssertValue (id, c)) } + | test_lpar_id_colon; "("; id = ident_or_anti; ":"; c = Constr.lconstr; ")"; tac = by_tactic -> + { let ipat = CAst.make ~loc @@ QIntroNaming (CAst.make ~loc @@ QIntroIdentifier id) in + CAst.make ~loc (QAssertType (Some ipat, c, tac)) } + | c = Constr.constr; ipat = as_ipat; tac = by_tactic -> + { CAst.make ~loc (QAssertType (ipat, c, tac)) } + ] ] + ; + q_assert: + [ [ a = assertion -> { a } ] ] + ; +END + +(** Extension of constr syntax *) + +(* +GRAMMAR EXTEND Gram + Pcoq.Constr.operconstr: LEVEL "0" + [ [ IDENT "ltac2"; ":"; "("; tac = tac2expr; ")" -> + { let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in + CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) } + | test_ampersand_ident; "&"; id = Prim.ident -> + { let tac = Tac2quote.of_exact_hyp ~loc (CAst.make ~loc id) in + let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in + CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) } + | test_dollar_ident; "$"; id = Prim.ident -> + { let id = Loc.tag ~loc id in + let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2_quotation) id in + CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) } + ] ] + ; +END +*) +{ + +let () = + +let open Extend in +let open Tok in +let (++) r s = Next (r, s) in +let rules = [ + Rule ( + Stop ++ Aentry test_dollar_ident ++ Atoken (PKEYWORD "$") ++ Aentry Prim.ident, + begin fun id _ _ loc -> + let id = Loc.tag ~loc id in + let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2_quotation) id in + CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) + end + ); + + Rule ( + Stop ++ Aentry test_ampersand_ident ++ Atoken (PKEYWORD "&") ++ Aentry Prim.ident, + begin fun id _ _ loc -> + let tac = Tac2quote.of_exact_hyp ~loc (CAst.make ~loc id) in + let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in + CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) + end + ); + + Rule ( + Stop ++ Atoken (PIDENT (Some "ltac2")) ++ Atoken (PKEYWORD ":") ++ + Atoken (PKEYWORD "(") ++ Aentry tac2expr ++ Atoken (PKEYWORD ")"), + begin fun _ tac _ _ _ loc -> + let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in + CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) + end + ) +] in + +Hook.set Tac2entries.register_constr_quotations begin fun () -> + Pcoq.grammar_extend Pcoq.Constr.operconstr None (Some (Gramlib.Gramext.Level "0"), [(None, None, rules)]) +end + +} + +{ + +let pr_ltac2entry _ = mt () (* FIXME *) +let pr_ltac2expr _ = mt () (* FIXME *) + +} + +VERNAC ARGUMENT EXTEND ltac2_entry +PRINTED BY { pr_ltac2entry } +| [ tac2def_val(v) ] -> { v } +| [ tac2def_typ(t) ] -> { t } +| [ tac2def_ext(e) ] -> { e } +| [ tac2def_syn(e) ] -> { e } +| [ tac2def_mut(e) ] -> { e } +| [ tac2def_run(e) ] -> { e } +END + +{ + +let classify_ltac2 = function +| StrSyn _ -> Vernacextend.(VtSideff [], VtNow) +| StrMut _ | StrVal _ | StrPrm _ | StrTyp _ | StrRun _ -> Vernacextend.classify_as_sideeff + +} + +VERNAC COMMAND EXTEND VernacDeclareTactic2Definition +| #[ local = locality ] ![proof] [ "Ltac2" ltac2_entry(e) ] => { classify_ltac2 e } -> { + fun ~pstate -> Tac2entries.register_struct ?local ~pstate e; pstate + } +END + +{ + +let _ = Pvernac.register_proof_mode "Ltac2" tac2mode + +} + +VERNAC ARGUMENT EXTEND ltac2_expr +PRINTED BY { pr_ltac2expr } +| [ tac2expr(e) ] -> { e } +END + +{ + +open G_ltac +open Vernacextend + +} + +VERNAC { tac2mode } EXTEND VernacLtac2 +| ![proof] [ ltac2_expr(t) ltac_use_default(default) ] => + { classify_as_proofstep } -> { +(* let g = Option.default (Proof_global.get_default_goal_selector ()) g in *) + fun ~pstate -> + Option.map (fun pstate -> Tac2entries.call ~pstate ~default t) pstate + } +END + +{ + +open Stdarg + +} + +VERNAC COMMAND EXTEND Ltac2Print CLASSIFIED AS SIDEFF +| [ "Print" "Ltac2" reference(tac) ] -> { Tac2entries.print_ltac tac } +END diff --git a/vendor/Ltac2/src/ltac2_plugin.mlpack b/vendor/Ltac2/src/ltac2_plugin.mlpack new file mode 100644 index 0000000000..2a25e825cb --- /dev/null +++ b/vendor/Ltac2/src/ltac2_plugin.mlpack @@ -0,0 +1,14 @@ +Tac2dyn +Tac2ffi +Tac2env +Tac2print +Tac2intern +Tac2interp +Tac2entries +Tac2quote +Tac2match +Tac2core +Tac2extffi +Tac2tactics +Tac2stdlib +G_ltac2 diff --git a/vendor/Ltac2/src/tac2core.ml b/vendor/Ltac2/src/tac2core.ml new file mode 100644 index 0000000000..d7e7b91ee6 --- /dev/null +++ b/vendor/Ltac2/src/tac2core.ml @@ -0,0 +1,1446 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Value.of_option Value.of_ident None +| Name id -> Value.of_option Value.of_ident (Some id) + +let to_name c = match Value.to_option Value.to_ident c with +| None -> Anonymous +| Some id -> Name id + +let of_relevance = function + | Sorts.Relevant -> ValInt 0 + | Sorts.Irrelevant -> ValInt 1 + +let to_relevance = function + | ValInt 0 -> Sorts.Relevant + | ValInt 1 -> Sorts.Irrelevant + | _ -> assert false + +let of_annot f Context.{binder_name;binder_relevance} = + of_tuple [|(f binder_name); of_relevance binder_relevance|] + +let to_annot f x = + match to_tuple x with + | [|x;y|] -> + let x = f x in + let y = to_relevance y in + Context.make_annot x y + | _ -> assert false + +let of_instance u = + let u = Univ.Instance.to_array (EConstr.Unsafe.to_instance u) in + Value.of_array (fun v -> Value.of_ext Value.val_univ v) u + +let to_instance u = + let u = Value.to_array (fun v -> Value.to_ext Value.val_univ v) u in + EConstr.EInstance.make (Univ.Instance.of_array u) + +let of_rec_declaration (nas, ts, cs) = + (Value.of_array (of_annot of_name) nas, + Value.of_array Value.of_constr ts, + Value.of_array Value.of_constr cs) + +let to_rec_declaration (nas, ts, cs) = + (Value.to_array (to_annot to_name) nas, + Value.to_array Value.to_constr ts, + Value.to_array Value.to_constr cs) + +let of_result f = function +| Inl c -> v_blk 0 [|f c|] +| Inr e -> v_blk 1 [|Value.of_exn e|] + +(** Stdlib exceptions *) + +let err_notfocussed = + Tac2interp.LtacError (coq_core "Not_focussed", [||]) + +let err_outofbounds = + Tac2interp.LtacError (coq_core "Out_of_bounds", [||]) + +let err_notfound = + Tac2interp.LtacError (coq_core "Not_found", [||]) + +let err_matchfailure = + Tac2interp.LtacError (coq_core "Match_failure", [||]) + +(** Helper functions *) + +let thaw f = Tac2ffi.apply f [v_unit] + +let fatal_flag : unit Exninfo.t = Exninfo.make () + +let set_bt info = + if !Tac2interp.print_ltac2_backtrace then + Tac2interp.get_backtrace >>= fun bt -> + Proofview.tclUNIT (Exninfo.add info Tac2entries.backtrace bt) + else Proofview.tclUNIT info + +let throw ?(info = Exninfo.null) e = + set_bt info >>= fun info -> + let info = Exninfo.add info fatal_flag () in + Proofview.tclLIFT (Proofview.NonLogical.raise ~info e) + +let fail ?(info = Exninfo.null) e = + set_bt info >>= fun info -> + Proofview.tclZERO ~info e + +let return x = Proofview.tclUNIT x +let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } + +let wrap f = + return () >>= fun () -> return (f ()) + +let wrap_unit f = + return () >>= fun () -> f (); return v_unit + +let assert_focussed = + Proofview.Goal.goals >>= fun gls -> + match gls with + | [_] -> Proofview.tclUNIT () + | [] | _ :: _ :: _ -> throw err_notfocussed + +let pf_apply f = + Proofview.Goal.goals >>= function + | [] -> + Proofview.tclENV >>= fun env -> + Proofview.tclEVARMAP >>= fun sigma -> + f env sigma + | [gl] -> + gl >>= fun gl -> + f (Proofview.Goal.env gl) (Tacmach.New.project gl) + | _ :: _ :: _ -> + throw err_notfocussed + +(** Primitives *) + +let define_primitive name arity f = + Tac2env.define_primitive (pname name) (mk_closure arity f) + +let define0 name f = define_primitive name arity_one (fun _ -> f) + +let define1 name r0 f = define_primitive name arity_one begin fun x -> + f (Value.repr_to r0 x) +end + +let define2 name r0 r1 f = define_primitive name (arity_suc arity_one) begin fun x y -> + f (Value.repr_to r0 x) (Value.repr_to r1 y) +end + +let define3 name r0 r1 r2 f = define_primitive name (arity_suc (arity_suc arity_one)) begin fun x y z -> + f (Value.repr_to r0 x) (Value.repr_to r1 y) (Value.repr_to r2 z) +end + +(** Printing *) + +let () = define1 "print" pp begin fun pp -> + wrap_unit (fun () -> Feedback.msg_notice pp) +end + +let () = define1 "message_of_int" int begin fun n -> + return (Value.of_pp (Pp.int n)) +end + +let () = define1 "message_of_string" string begin fun s -> + return (Value.of_pp (str (Bytes.to_string s))) +end + +let () = define1 "message_of_constr" constr begin fun c -> + pf_apply begin fun env sigma -> + let pp = Printer.pr_econstr_env env sigma c in + return (Value.of_pp pp) + end +end + +let () = define1 "message_of_ident" ident begin fun c -> + let pp = Id.print c in + return (Value.of_pp pp) +end + +let () = define1 "message_of_exn" valexpr begin fun v -> + Proofview.tclENV >>= fun env -> + Proofview.tclEVARMAP >>= fun sigma -> + let pp = Tac2print.pr_valexpr env sigma v (GTypRef (Other Core.t_exn, [])) in + return (Value.of_pp pp) +end + + +let () = define2 "message_concat" pp pp begin fun m1 m2 -> + return (Value.of_pp (Pp.app m1 m2)) +end + +(** Array *) + +let () = define2 "array_make" int valexpr begin fun n x -> + if n < 0 || n > Sys.max_array_length then throw err_outofbounds + else wrap (fun () -> v_blk 0 (Array.make n x)) +end + +let () = define1 "array_length" block begin fun (_, v) -> + return (Value.of_int (Array.length v)) +end + +let () = define3 "array_set" block int valexpr begin fun (_, v) n x -> + if n < 0 || n >= Array.length v then throw err_outofbounds + else wrap_unit (fun () -> v.(n) <- x) +end + +let () = define2 "array_get" block int begin fun (_, v) n -> + if n < 0 || n >= Array.length v then throw err_outofbounds + else wrap (fun () -> v.(n)) +end + +(** Ident *) + +let () = define2 "ident_equal" ident ident begin fun id1 id2 -> + return (Value.of_bool (Id.equal id1 id2)) +end + +let () = define1 "ident_to_string" ident begin fun id -> + return (Value.of_string (Bytes.of_string (Id.to_string id))) +end + +let () = define1 "ident_of_string" string begin fun s -> + let id = try Some (Id.of_string (Bytes.to_string s)) with _ -> None in + return (Value.of_option Value.of_ident id) +end + +(** Int *) + +let () = define2 "int_equal" int int begin fun m n -> + return (Value.of_bool (m == n)) +end + +let binop n f = define2 n int int begin fun m n -> + return (Value.of_int (f m n)) +end + +let () = binop "int_compare" Int.compare +let () = binop "int_add" (+) +let () = binop "int_sub" (-) +let () = binop "int_mul" ( * ) + +let () = define1 "int_neg" int begin fun m -> + return (Value.of_int (~- m)) +end + +(** Char *) + +let () = define1 "char_of_int" int begin fun n -> + wrap (fun () -> Value.of_char (Char.chr n)) +end + +let () = define1 "char_to_int" char begin fun n -> + wrap (fun () -> Value.of_int (Char.code n)) +end + +(** String *) + +let () = define2 "string_make" int char begin fun n c -> + if n < 0 || n > Sys.max_string_length then throw err_outofbounds + else wrap (fun () -> Value.of_string (Bytes.make n c)) +end + +let () = define1 "string_length" string begin fun s -> + return (Value.of_int (Bytes.length s)) +end + +let () = define3 "string_set" string int char begin fun s n c -> + if n < 0 || n >= Bytes.length s then throw err_outofbounds + else wrap_unit (fun () -> Bytes.set s n c) +end + +let () = define2 "string_get" string int begin fun s n -> + if n < 0 || n >= Bytes.length s then throw err_outofbounds + else wrap (fun () -> Value.of_char (Bytes.get s n)) +end + +(** Terms *) + +(** constr -> constr *) +let () = define1 "constr_type" constr begin fun c -> + let get_type env sigma = + Proofview.V82.wrap_exceptions begin fun () -> + let (sigma, t) = Typing.type_of env sigma c in + let t = Value.of_constr t in + Proofview.Unsafe.tclEVARS sigma <*> Proofview.tclUNIT t + end in + pf_apply get_type +end + +(** constr -> constr *) +let () = define2 "constr_equal" constr constr begin fun c1 c2 -> + Proofview.tclEVARMAP >>= fun sigma -> + let b = EConstr.eq_constr sigma c1 c2 in + Proofview.tclUNIT (Value.of_bool b) +end + +let () = define1 "constr_kind" constr begin fun c -> + let open Constr in + Proofview.tclEVARMAP >>= fun sigma -> + return begin match EConstr.kind sigma c with + | Rel n -> + v_blk 0 [|Value.of_int n|] + | Var id -> + v_blk 1 [|Value.of_ident id|] + | Meta n -> + v_blk 2 [|Value.of_int n|] + | Evar (evk, args) -> + v_blk 3 [| + Value.of_int (Evar.repr evk); + Value.of_array Value.of_constr args; + |] + | Sort s -> + v_blk 4 [|Value.of_ext Value.val_sort s|] + | Cast (c, k, t) -> + v_blk 5 [| + Value.of_constr c; + Value.of_ext Value.val_cast k; + Value.of_constr t; + |] + | Prod (na, t, u) -> + v_blk 6 [| + of_annot of_name na; + Value.of_constr t; + Value.of_constr u; + |] + | Lambda (na, t, c) -> + v_blk 7 [| + of_annot of_name na; + Value.of_constr t; + Value.of_constr c; + |] + | LetIn (na, b, t, c) -> + v_blk 8 [| + of_annot of_name na; + Value.of_constr b; + Value.of_constr t; + Value.of_constr c; + |] + | App (c, cl) -> + v_blk 9 [| + Value.of_constr c; + Value.of_array Value.of_constr cl; + |] + | Const (cst, u) -> + v_blk 10 [| + Value.of_constant cst; + of_instance u; + |] + | Ind (ind, u) -> + v_blk 11 [| + Value.of_ext Value.val_inductive ind; + of_instance u; + |] + | Construct (cstr, u) -> + v_blk 12 [| + Value.of_ext Value.val_constructor cstr; + of_instance u; + |] + | Case (ci, c, t, bl) -> + v_blk 13 [| + Value.of_ext Value.val_case ci; + Value.of_constr c; + Value.of_constr t; + Value.of_array Value.of_constr bl; + |] + | Fix ((recs, i), def) -> + let (nas, ts, cs) = of_rec_declaration def in + v_blk 14 [| + Value.of_array Value.of_int recs; + Value.of_int i; + nas; + ts; + cs; + |] + | CoFix (i, def) -> + let (nas, ts, cs) = of_rec_declaration def in + v_blk 15 [| + Value.of_int i; + nas; + ts; + cs; + |] + | Proj (p, c) -> + v_blk 16 [| + Value.of_ext Value.val_projection p; + Value.of_constr c; + |] + | Int _ -> + assert false + end +end + +let () = define1 "constr_make" valexpr begin fun knd -> + let c = match Tac2ffi.to_block knd with + | (0, [|n|]) -> + let n = Value.to_int n in + EConstr.mkRel n + | (1, [|id|]) -> + let id = Value.to_ident id in + EConstr.mkVar id + | (2, [|n|]) -> + let n = Value.to_int n in + EConstr.mkMeta n + | (3, [|evk; args|]) -> + let evk = Evar.unsafe_of_int (Value.to_int evk) in + let args = Value.to_array Value.to_constr args in + EConstr.mkEvar (evk, args) + | (4, [|s|]) -> + let s = Value.to_ext Value.val_sort s in + EConstr.mkSort (EConstr.Unsafe.to_sorts s) + | (5, [|c; k; t|]) -> + let c = Value.to_constr c in + let k = Value.to_ext Value.val_cast k in + let t = Value.to_constr t in + EConstr.mkCast (c, k, t) + | (6, [|na; t; u|]) -> + let na = to_annot to_name na in + let t = Value.to_constr t in + let u = Value.to_constr u in + EConstr.mkProd (na, t, u) + | (7, [|na; t; c|]) -> + let na = to_annot to_name na in + let t = Value.to_constr t in + let u = Value.to_constr c in + EConstr.mkLambda (na, t, u) + | (8, [|na; b; t; c|]) -> + let na = to_annot to_name na in + let b = Value.to_constr b in + let t = Value.to_constr t in + let c = Value.to_constr c in + EConstr.mkLetIn (na, b, t, c) + | (9, [|c; cl|]) -> + let c = Value.to_constr c in + let cl = Value.to_array Value.to_constr cl in + EConstr.mkApp (c, cl) + | (10, [|cst; u|]) -> + let cst = Value.to_constant cst in + let u = to_instance u in + EConstr.mkConstU (cst, u) + | (11, [|ind; u|]) -> + let ind = Value.to_ext Value.val_inductive ind in + let u = to_instance u in + EConstr.mkIndU (ind, u) + | (12, [|cstr; u|]) -> + let cstr = Value.to_ext Value.val_constructor cstr in + let u = to_instance u in + EConstr.mkConstructU (cstr, u) + | (13, [|ci; c; t; bl|]) -> + let ci = Value.to_ext Value.val_case ci in + let c = Value.to_constr c in + let t = Value.to_constr t in + let bl = Value.to_array Value.to_constr bl in + EConstr.mkCase (ci, c, t, bl) + | (14, [|recs; i; nas; ts; cs|]) -> + let recs = Value.to_array Value.to_int recs in + let i = Value.to_int i in + let def = to_rec_declaration (nas, ts, cs) in + EConstr.mkFix ((recs, i), def) + | (15, [|i; nas; ts; cs|]) -> + let i = Value.to_int i in + let def = to_rec_declaration (nas, ts, cs) in + EConstr.mkCoFix (i, def) + | (16, [|p; c|]) -> + let p = Value.to_ext Value.val_projection p in + let c = Value.to_constr c in + EConstr.mkProj (p, c) + | _ -> assert false + in + return (Value.of_constr c) +end + +let () = define1 "constr_check" constr begin fun c -> + pf_apply begin fun env sigma -> + try + let (sigma, _) = Typing.type_of env sigma c in + Proofview.Unsafe.tclEVARS sigma >>= fun () -> + return (of_result Value.of_constr (Inl c)) + with e when CErrors.noncritical e -> + let e = CErrors.push e in + return (of_result Value.of_constr (Inr e)) + end +end + +let () = define3 "constr_substnl" (list constr) int constr begin fun subst k c -> + let ans = EConstr.Vars.substnl subst k c in + return (Value.of_constr ans) +end + +let () = define3 "constr_closenl" (list ident) int constr begin fun ids k c -> + let ans = EConstr.Vars.substn_vars k ids c in + return (Value.of_constr ans) +end + +let () = define1 "constr_case" (repr_ext val_inductive) begin fun ind -> + Proofview.tclENV >>= fun env -> + try + let ans = Inductiveops.make_case_info env ind Sorts.Relevant Constr.RegularStyle in + return (Value.of_ext Value.val_case ans) + with e when CErrors.noncritical e -> + throw err_notfound +end + +let () = define2 "constr_constructor" (repr_ext val_inductive) int begin fun (ind, i) k -> + Proofview.tclENV >>= fun env -> + try + let open Declarations in + let ans = Environ.lookup_mind ind env in + let _ = ans.mind_packets.(i).mind_consnames.(k) in + return (Value.of_ext val_constructor ((ind, i), (k + 1))) + with e when CErrors.noncritical e -> + throw err_notfound +end + +let () = define3 "constr_in_context" ident constr closure begin fun id t c -> + Proofview.Goal.goals >>= function + | [gl] -> + gl >>= fun gl -> + let env = Proofview.Goal.env gl in + let sigma = Proofview.Goal.sigma gl in + let has_var = + try + let _ = Environ.lookup_named_val id env in + true + with Not_found -> false + in + if has_var then + Tacticals.New.tclZEROMSG (str "Variable already exists") + else + let open Context.Named.Declaration in + let nenv = EConstr.push_named (LocalAssum (Context.make_annot id Sorts.Relevant, t)) env in + let (sigma, (evt, _)) = Evarutil.new_type_evar nenv sigma Evd.univ_flexible in + let (sigma, evk) = Evarutil.new_pure_evar (Environ.named_context_val nenv) sigma evt in + Proofview.Unsafe.tclEVARS sigma >>= fun () -> + Proofview.Unsafe.tclSETGOALS [Proofview.with_empty_state evk] >>= fun () -> + thaw c >>= fun _ -> + Proofview.Unsafe.tclSETGOALS [Proofview.with_empty_state (Proofview.Goal.goal gl)] >>= fun () -> + let args = List.map (fun d -> EConstr.mkVar (get_id d)) (EConstr.named_context env) in + let args = Array.of_list (EConstr.mkRel 1 :: args) in + let ans = EConstr.mkEvar (evk, args) in + let ans = EConstr.mkLambda (Context.make_annot (Name id) Sorts.Relevant, t, ans) in + return (Value.of_constr ans) + | _ -> + throw err_notfocussed +end + +(** Patterns *) + +let empty_context = EConstr.mkMeta Constr_matching.special_meta + +let () = define0 "pattern_empty_context" begin + return (Value.of_constr empty_context) +end + +let () = define2 "pattern_matches" pattern constr begin fun pat c -> + pf_apply begin fun env sigma -> + let ans = + try Some (Constr_matching.matches env sigma pat c) + with Constr_matching.PatternMatchingFailure -> None + in + begin match ans with + | None -> fail err_matchfailure + | Some ans -> + let ans = Id.Map.bindings ans in + let of_pair (id, c) = Value.of_tuple [| Value.of_ident id; Value.of_constr c |] in + return (Value.of_list of_pair ans) + end + end +end + +let () = define2 "pattern_matches_subterm" pattern constr begin fun pat c -> + let open Constr_matching in + let rec of_ans s = match IStream.peek s with + | IStream.Nil -> fail err_matchfailure + | IStream.Cons ({ m_sub = (_, sub); m_ctx }, s) -> + let ans = Id.Map.bindings sub in + let of_pair (id, c) = Value.of_tuple [| Value.of_ident id; Value.of_constr c |] in + let ans = Value.of_tuple [| Value.of_constr (Lazy.force m_ctx); Value.of_list of_pair ans |] in + Proofview.tclOR (return ans) (fun _ -> of_ans s) + in + pf_apply begin fun env sigma -> + let ans = Constr_matching.match_subterm env sigma (Id.Set.empty,pat) c in + of_ans ans + end +end + +let () = define2 "pattern_matches_vect" pattern constr begin fun pat c -> + pf_apply begin fun env sigma -> + let ans = + try Some (Constr_matching.matches env sigma pat c) + with Constr_matching.PatternMatchingFailure -> None + in + begin match ans with + | None -> fail err_matchfailure + | Some ans -> + let ans = Id.Map.bindings ans in + let ans = Array.map_of_list snd ans in + return (Value.of_array Value.of_constr ans) + end + end +end + +let () = define2 "pattern_matches_subterm_vect" pattern constr begin fun pat c -> + let open Constr_matching in + let rec of_ans s = match IStream.peek s with + | IStream.Nil -> fail err_matchfailure + | IStream.Cons ({ m_sub = (_, sub); m_ctx }, s) -> + let ans = Id.Map.bindings sub in + let ans = Array.map_of_list snd ans in + let ans = Value.of_tuple [| Value.of_constr (Lazy.force m_ctx); Value.of_array Value.of_constr ans |] in + Proofview.tclOR (return ans) (fun _ -> of_ans s) + in + pf_apply begin fun env sigma -> + let ans = Constr_matching.match_subterm env sigma (Id.Set.empty,pat) c in + of_ans ans + end +end + +let () = define3 "pattern_matches_goal" bool (list (pair bool pattern)) (pair bool pattern) begin fun rev hp cp -> + assert_focussed >>= fun () -> + Proofview.Goal.enter_one begin fun gl -> + let env = Proofview.Goal.env gl in + let sigma = Proofview.Goal.sigma gl in + let concl = Proofview.Goal.concl gl in + let mk_pattern (b, pat) = if b then Tac2match.MatchPattern pat else Tac2match.MatchContext pat in + let r = (List.map mk_pattern hp, mk_pattern cp) in + Tac2match.match_goal env sigma concl ~rev r >>= fun (hyps, ctx, subst) -> + let of_ctxopt ctx = Value.of_constr (Option.default empty_context ctx) in + let hids = Value.of_array Value.of_ident (Array.map_of_list fst hyps) in + let hctx = Value.of_array of_ctxopt (Array.map_of_list snd hyps) in + let subs = Value.of_array Value.of_constr (Array.map_of_list snd (Id.Map.bindings subst)) in + let cctx = of_ctxopt ctx in + let ans = Value.of_tuple [| hids; hctx; subs; cctx |] in + Proofview.tclUNIT ans + end +end + +let () = define2 "pattern_instantiate" constr constr begin fun ctx c -> + let ctx = EConstr.Unsafe.to_constr ctx in + let c = EConstr.Unsafe.to_constr c in + let ans = Termops.subst_meta [Constr_matching.special_meta, c] ctx in + return (Value.of_constr (EConstr.of_constr ans)) +end + +(** Error *) + +let () = define1 "throw" exn begin fun (e, info) -> + throw ~info e +end + +(** Control *) + +(** exn -> 'a *) +let () = define1 "zero" exn begin fun (e, info) -> + fail ~info e +end + +(** (unit -> 'a) -> (exn -> 'a) -> 'a *) +let () = define2 "plus" closure closure begin fun x k -> + Proofview.tclOR (thaw x) (fun e -> Tac2ffi.apply k [Value.of_exn e]) +end + +(** (unit -> 'a) -> 'a *) +let () = define1 "once" closure begin fun f -> + Proofview.tclONCE (thaw f) +end + +(** (unit -> unit) list -> unit *) +let () = define1 "dispatch" (list closure) begin fun l -> + let l = List.map (fun f -> Proofview.tclIGNORE (thaw f)) l in + Proofview.tclDISPATCH l >>= fun () -> return v_unit +end + +(** (unit -> unit) list -> (unit -> unit) -> (unit -> unit) list -> unit *) +let () = define3 "extend" (list closure) closure (list closure) begin fun lft tac rgt -> + let lft = List.map (fun f -> Proofview.tclIGNORE (thaw f)) lft in + let tac = Proofview.tclIGNORE (thaw tac) in + let rgt = List.map (fun f -> Proofview.tclIGNORE (thaw f)) rgt in + Proofview.tclEXTEND lft tac rgt >>= fun () -> return v_unit +end + +(** (unit -> unit) -> unit *) +let () = define1 "enter" closure begin fun f -> + let f = Proofview.tclIGNORE (thaw f) in + Proofview.tclINDEPENDENT f >>= fun () -> return v_unit +end + +(** (unit -> 'a) -> ('a * ('exn -> 'a)) result *) +let () = define1 "case" closure begin fun f -> + Proofview.tclCASE (thaw f) >>= begin function + | Proofview.Next (x, k) -> + let k = Tac2ffi.mk_closure arity_one begin fun e -> + let (e, info) = Value.to_exn e in + set_bt info >>= fun info -> + k (e, info) + end in + return (v_blk 0 [| Value.of_tuple [| x; Value.of_closure k |] |]) + | Proofview.Fail e -> return (v_blk 1 [| Value.of_exn e |]) + end +end + +(** int -> int -> (unit -> 'a) -> 'a *) +let () = define3 "focus" int int closure begin fun i j tac -> + Proofview.tclFOCUS i j (thaw tac) +end + +(** unit -> unit *) +let () = define0 "shelve" begin + Proofview.shelve >>= fun () -> return v_unit +end + +(** unit -> unit *) +let () = define0 "shelve_unifiable" begin + Proofview.shelve_unifiable >>= fun () -> return v_unit +end + +let () = define1 "new_goal" int begin fun ev -> + let ev = Evar.unsafe_of_int ev in + Proofview.tclEVARMAP >>= fun sigma -> + if Evd.mem sigma ev then + Proofview.Unsafe.tclNEWGOALS [Proofview.with_empty_state ev] <*> Proofview.tclUNIT v_unit + else throw err_notfound +end + +(** unit -> constr *) +let () = define0 "goal" begin + assert_focussed >>= fun () -> + Proofview.Goal.enter_one begin fun gl -> + let concl = Tacmach.New.pf_nf_concl gl in + return (Value.of_constr concl) + end +end + +(** ident -> constr *) +let () = define1 "hyp" ident begin fun id -> + pf_apply begin fun env _ -> + let mem = try ignore (Environ.lookup_named id env); true with Not_found -> false in + if mem then return (Value.of_constr (EConstr.mkVar id)) + else Tacticals.New.tclZEROMSG + (str "Hypothesis " ++ quote (Id.print id) ++ str " not found") (* FIXME: Do something more sensible *) + end +end + +let () = define0 "hyps" begin + pf_apply begin fun env _ -> + let open Context in + let open Named.Declaration in + let hyps = List.rev (Environ.named_context env) in + let map = function + | LocalAssum (id, t) -> + let t = EConstr.of_constr t in + Value.of_tuple [|Value.of_ident id.binder_name; Value.of_option Value.of_constr None; Value.of_constr t|] + | LocalDef (id, c, t) -> + let c = EConstr.of_constr c in + let t = EConstr.of_constr t in + Value.of_tuple [|Value.of_ident id.binder_name; Value.of_option Value.of_constr (Some c); Value.of_constr t|] + in + return (Value.of_list map hyps) + end +end + +(** (unit -> constr) -> unit *) +let () = define1 "refine" closure begin fun c -> + let c = thaw c >>= fun c -> Proofview.tclUNIT ((), Value.to_constr c) in + Proofview.Goal.enter begin fun gl -> + Refine.generic_refine ~typecheck:true c gl + end >>= fun () -> return v_unit +end + +let () = define2 "with_holes" closure closure begin fun x f -> + Proofview.tclEVARMAP >>= fun sigma0 -> + thaw x >>= fun ans -> + Proofview.tclEVARMAP >>= fun sigma -> + Proofview.Unsafe.tclEVARS sigma0 >>= fun () -> + Tacticals.New.tclWITHHOLES false (Tac2ffi.apply f [ans]) sigma +end + +let () = define1 "progress" closure begin fun f -> + Proofview.tclPROGRESS (thaw f) +end + +let () = define2 "abstract" (option ident) closure begin fun id f -> + Abstract.tclABSTRACT id (Proofview.tclIGNORE (thaw f)) >>= fun () -> + return v_unit +end + +let () = define2 "time" (option string) closure begin fun s f -> + let s = Option.map Bytes.to_string s in + Proofview.tclTIME s (thaw f) +end + +let () = define0 "check_interrupt" begin + Proofview.tclCHECKINTERRUPT >>= fun () -> return v_unit +end + +(** Fresh *) + +let () = define2 "fresh_free_union" (repr_ext val_free) (repr_ext val_free) begin fun set1 set2 -> + let ans = Id.Set.union set1 set2 in + return (Value.of_ext Value.val_free ans) +end + +let () = define1 "fresh_free_of_ids" (list ident) begin fun ids -> + let free = List.fold_right Id.Set.add ids Id.Set.empty in + return (Value.of_ext Value.val_free free) +end + +let () = define1 "fresh_free_of_constr" constr begin fun c -> + Proofview.tclEVARMAP >>= fun sigma -> + let rec fold accu c = match EConstr.kind sigma c with + | Constr.Var id -> Id.Set.add id accu + | _ -> EConstr.fold sigma fold accu c + in + let ans = fold Id.Set.empty c in + return (Value.of_ext Value.val_free ans) +end + +let () = define2 "fresh_fresh" (repr_ext val_free) ident begin fun avoid id -> + let nid = Namegen.next_ident_away_from id (fun id -> Id.Set.mem id avoid) in + return (Value.of_ident nid) +end + +(** Env *) + +let () = define1 "env_get" (list ident) begin fun ids -> + let r = match ids with + | [] -> None + | _ :: _ as ids -> + let (id, path) = List.sep_last ids in + let path = DirPath.make (List.rev path) in + let fp = Libnames.make_path path id in + try Some (Nametab.global_of_path fp) with Not_found -> None + in + return (Value.of_option Value.of_reference r) +end + +let () = define1 "env_expand" (list ident) begin fun ids -> + let r = match ids with + | [] -> [] + | _ :: _ as ids -> + let (id, path) = List.sep_last ids in + let path = DirPath.make (List.rev path) in + let qid = Libnames.make_qualid path id in + Nametab.locate_all qid + in + return (Value.of_list Value.of_reference r) +end + +let () = define1 "env_path" reference begin fun r -> + match Nametab.path_of_global r with + | fp -> + let (path, id) = Libnames.repr_path fp in + let path = DirPath.repr path in + return (Value.of_list Value.of_ident (List.rev_append path [id])) + | exception Not_found -> + throw err_notfound +end + +let () = define1 "env_instantiate" reference begin fun r -> + Proofview.tclENV >>= fun env -> + Proofview.tclEVARMAP >>= fun sigma -> + let (sigma, c) = Evd.fresh_global env sigma r in + Proofview.Unsafe.tclEVARS sigma >>= fun () -> + return (Value.of_constr c) +end + +(** Ltac1 in Ltac2 *) + +let ltac1 = Tac2ffi.repr_ext Value.val_ltac1 +let of_ltac1 v = Value.of_ext Value.val_ltac1 v + +let () = define1 "ltac1_ref" (list ident) begin fun ids -> + let open Ltac_plugin in + let r = match ids with + | [] -> raise Not_found + | _ :: _ as ids -> + let (id, path) = List.sep_last ids in + let path = DirPath.make (List.rev path) in + let fp = Libnames.make_path path id in + if Tacenv.exists_tactic fp then + List.hd (Tacenv.locate_extended_all_tactic (Libnames.qualid_of_path fp)) + else raise Not_found + in + let tac = Tacinterp.Value.of_closure (Tacinterp.default_ist ()) (Tacenv.interp_ltac r) in + return (Value.of_ext val_ltac1 tac) +end + +let () = define1 "ltac1_run" ltac1 begin fun v -> + let open Ltac_plugin in + Tacinterp.tactic_of_value (Tacinterp.default_ist ()) v >>= fun () -> + return v_unit +end + +let () = define3 "ltac1_apply" ltac1 (list ltac1) closure begin fun f args k -> + let open Ltac_plugin in + let open Tacexpr in + let open Locus in + let k ret = + Proofview.tclIGNORE (Tac2ffi.apply k [Value.of_ext val_ltac1 ret]) + in + let fold arg (i, vars, lfun) = + let id = Id.of_string ("x" ^ string_of_int i) in + let x = Reference (ArgVar CAst.(make 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 + let tac = TacArg(CAst.make @@ TacCall (CAst.make (ArgVar CAst.(make @@ Id.of_string "F"),args))) in + Tacinterp.val_interp ist tac k >>= fun () -> + return v_unit +end + +let () = define1 "ltac1_of_constr" constr begin fun c -> + let open Ltac_plugin in + return (Value.of_ext val_ltac1 (Tacinterp.Value.of_constr c)) +end + +let () = define1 "ltac1_to_constr" ltac1 begin fun v -> + let open Ltac_plugin in + return (Value.of_option Value.of_constr (Tacinterp.Value.to_constr v)) +end + +let () = define1 "ltac1_of_list" (list ltac1) begin fun l -> + let open Geninterp.Val in + return (Value.of_ext val_ltac1 (inject (Base typ_list) l)) +end + +let () = define1 "ltac1_to_list" ltac1 begin fun v -> + let open Ltac_plugin in + return (Value.of_option (Value.of_list of_ltac1) (Tacinterp.Value.to_list v)) +end + +(** ML types *) + +let constr_flags () = + let open Pretyping in + { + use_typeclasses = true; + solve_unification_constraints = true; + fail_evar = true; + expand_evars = true; + program_mode = false; + polymorphic = false; + } + +let open_constr_no_classes_flags () = + let open Pretyping in + { + use_typeclasses = false; + solve_unification_constraints = true; + fail_evar = false; + expand_evars = true; + program_mode = false; + polymorphic = false; + } + +(** Embed all Ltac2 data into Values *) +let to_lvar ist = + let open Glob_ops in + let lfun = Tac2interp.set_env ist Id.Map.empty in + { empty_lvar with Ltac_pretype.ltac_genargs = lfun } + +let gtypref kn = GTypRef (Other kn, []) + +let intern_constr self ist c = + let (_, (c, _)) = Genintern.intern Stdarg.wit_constr ist c in + (GlbVal c, gtypref t_constr) + +let catchable_exception = function + | Logic_monad.Exception _ -> false + | e -> CErrors.noncritical e + +let interp_constr flags ist c = + let open Pretyping in + let ist = to_lvar ist in + pf_apply begin fun env sigma -> + try + let (sigma, c) = understand_ltac flags env sigma ist WithoutTypeConstraint c in + let c = Value.of_constr c in + Proofview.Unsafe.tclEVARS sigma >>= fun () -> + Proofview.tclUNIT c + with e when catchable_exception e -> + let (e, info) = CErrors.push e in + set_bt info >>= fun info -> + match Exninfo.get info fatal_flag with + | None -> Proofview.tclZERO ~info e + | Some () -> throw ~info e + end + +let () = + let intern = intern_constr in + let interp ist c = interp_constr (constr_flags ()) ist c in + let print env c = str "constr:(" ++ Printer.pr_lglob_constr_env env c ++ str ")" in + let subst subst c = Detyping.subst_glob_constr (Global.env()) subst c in + let obj = { + ml_intern = intern; + ml_subst = subst; + ml_interp = interp; + ml_print = print; + } in + define_ml_object Tac2quote.wit_constr obj + +let () = + let intern = intern_constr in + let interp ist c = interp_constr (open_constr_no_classes_flags ()) ist c in + let print env c = str "open_constr:(" ++ Printer.pr_lglob_constr_env env c ++ str ")" in + let subst subst c = Detyping.subst_glob_constr (Global.env()) subst c in + let obj = { + ml_intern = intern; + ml_subst = subst; + ml_interp = interp; + ml_print = print; + } in + define_ml_object Tac2quote.wit_open_constr obj + +let () = + let interp _ id = return (Value.of_ident id) in + let print _ id = str "ident:(" ++ Id.print id ++ str ")" in + let obj = { + ml_intern = (fun _ _ id -> GlbVal id, gtypref t_ident); + ml_interp = interp; + ml_subst = (fun _ id -> id); + ml_print = print; + } in + define_ml_object Tac2quote.wit_ident obj + +let () = + let intern self ist c = + let env = ist.Genintern.genv in + let sigma = Evd.from_env env in + let warn = if !Ltac_plugin.Tacintern.strict_check then fun x -> x else Constrintern.for_grammar in + let _, pat = warn (fun () ->Constrintern.intern_constr_pattern env sigma ~as_type:false c) () in + GlbVal pat, gtypref t_pattern + in + let subst subst c = + let env = Global.env () in + let sigma = Evd.from_env env in + Patternops.subst_pattern env sigma subst c + in + let print env pat = str "pattern:(" ++ Printer.pr_lconstr_pattern_env env Evd.empty pat ++ str ")" in + let interp _ c = return (Value.of_pattern c) in + let obj = { + ml_intern = intern; + ml_interp = interp; + ml_subst = subst; + ml_print = print; + } in + define_ml_object Tac2quote.wit_pattern obj + +let () = + let intern self ist ref = match ref.CAst.v with + | Tac2qexpr.QHypothesis id -> + GlbVal (Globnames.VarRef id), gtypref t_reference + | Tac2qexpr.QReference qid -> + let gr = + try Nametab.locate qid + with Not_found -> + Nametab.error_global_not_found qid + in + GlbVal gr, gtypref t_reference + in + let subst s c = Globnames.subst_global_reference s c in + let interp _ gr = return (Value.of_reference gr) in + let print _ = function + | Globnames.VarRef id -> str "reference:(" ++ str "&" ++ Id.print id ++ str ")" + | r -> str "reference:(" ++ Printer.pr_global r ++ str ")" + in + let obj = { + ml_intern = intern; + ml_subst = subst; + ml_interp = interp; + ml_print = print; + } in + define_ml_object Tac2quote.wit_reference obj + +let () = + let intern self ist tac = + (* Prevent inner calls to Ltac2 values *) + let extra = Tac2intern.drop_ltac2_env ist.Genintern.extra in + let ist = { ist with Genintern.extra } in + let _, tac = Genintern.intern Ltac_plugin.Tacarg.wit_tactic ist tac in + GlbVal tac, gtypref t_unit + in + let interp ist tac = + let ist = { env_ist = Id.Map.empty } in + let lfun = Tac2interp.set_env ist Id.Map.empty in + let ist = Ltac_plugin.Tacinterp.default_ist () in + let ist = { ist with Geninterp.lfun = lfun } in + let tac = (Ltac_plugin.Tacinterp.eval_tactic_ist ist tac : unit Proofview.tactic) in + let wrap (e, info) = set_bt info >>= fun info -> Proofview.tclZERO ~info e in + Proofview.tclOR tac wrap >>= fun () -> + return v_unit + in + let subst s tac = Genintern.substitute Ltac_plugin.Tacarg.wit_tactic s tac in + let print env tac = + str "ltac1:(" ++ Ltac_plugin.Pptactic.pr_glob_tactic env tac ++ str ")" + in + let obj = { + ml_intern = intern; + ml_subst = subst; + ml_interp = interp; + ml_print = print; + } in + define_ml_object Tac2quote.wit_ltac1 obj + +let () = + let open Ltac_plugin in + let intern self ist tac = + (* Prevent inner calls to Ltac2 values *) + let extra = Tac2intern.drop_ltac2_env ist.Genintern.extra in + let ist = { ist with Genintern.extra } in + let _, tac = Genintern.intern Ltac_plugin.Tacarg.wit_tactic ist tac in + GlbVal tac, gtypref t_ltac1 + in + let interp ist tac = + let ist = { env_ist = Id.Map.empty } in + let lfun = Tac2interp.set_env ist Id.Map.empty in + let ist = Ltac_plugin.Tacinterp.default_ist () in + let ist = { ist with Geninterp.lfun = lfun } in + return (Value.of_ext val_ltac1 (Tacinterp.Value.of_closure ist tac)) + in + let subst s tac = Genintern.substitute Tacarg.wit_tactic s tac in + let print env tac = + str "ltac1val:(" ++ Ltac_plugin.Pptactic.pr_glob_tactic env tac ++ str ")" + in + let obj = { + ml_intern = intern; + ml_subst = subst; + ml_interp = interp; + ml_print = print; + } in + define_ml_object Tac2quote.wit_ltac1val obj + +(** Ltac2 in terms *) + +let () = + let interp ist poly env sigma concl tac = + let ist = Tac2interp.get_env ist in + let tac = Proofview.tclIGNORE (Tac2interp.interp ist tac) in + let name, poly = Id.of_string "ltac2", poly in + let c, sigma = Pfedit.refine_by_tactic ~name ~poly env sigma concl tac in + (EConstr.of_constr c, sigma) + in + GlobEnv.register_constr_interp0 wit_ltac2 interp + +let () = + let interp ist poly env sigma concl id = + let ist = Tac2interp.get_env ist in + let c = Id.Map.find id ist.env_ist in + let c = Value.to_constr c in + let sigma = Typing.check env sigma c concl in + (c, sigma) + in + GlobEnv.register_constr_interp0 wit_ltac2_quotation interp + +let () = + let pr_raw id = Genprint.PrinterBasic (fun _env _sigma -> mt ()) in + let pr_glb id = Genprint.PrinterBasic (fun _env _sigma -> str "$" ++ Id.print id) in + let pr_top _ = Genprint.TopPrinterBasic mt in + Genprint.register_print0 wit_ltac2_quotation pr_raw pr_glb pr_top + +(** Ltac2 in Ltac1 *) + +let () = + let e = Tac2entries.Pltac.tac2expr in + let inject (loc, v) = Ltac_plugin.Tacexpr.TacGeneric (in_gen (rawwit wit_ltac2) v) in + Ltac_plugin.Tacentries.create_ltac_quotation "ltac2" inject (e, None) + +let () = + let open Ltac_plugin in + let open Tacinterp in + let idtac = Value.of_closure (default_ist ()) (Tacexpr.TacId []) in + let interp ist tac = +(* let ist = Tac2interp.get_env ist.Geninterp.lfun in *) + let ist = { env_ist = Id.Map.empty } in + Tac2interp.interp ist tac >>= fun _ -> + Ftactic.return idtac + in + Geninterp.register_interp0 wit_ltac2 interp + +let () = + let pr_raw _ = Genprint.PrinterBasic (fun _env _sigma -> mt ()) in + let pr_glb e = Genprint.PrinterBasic (fun _env _sigma -> Tac2print.pr_glbexpr e) in + let pr_top _ = Genprint.TopPrinterBasic mt in + Genprint.register_print0 wit_ltac2 pr_raw pr_glb pr_top + +(** Built-in notation scopes *) + +let add_scope s f = + Tac2entries.register_scope (Id.of_string s) f + +let rec pr_scope = let open CAst in function +| SexprStr {v=s} -> qstring s +| SexprInt {v=n} -> Pp.int n +| SexprRec (_, {v=na}, args) -> + let na = match na with + | None -> str "_" + | Some id -> Id.print id + in + na ++ str "(" ++ prlist_with_sep (fun () -> str ", ") pr_scope args ++ str ")" + +let scope_fail s args = + let args = str "(" ++ prlist_with_sep (fun () -> str ", ") pr_scope args ++ str ")" in + CErrors.user_err (str "Invalid arguments " ++ args ++ str " in scope " ++ str s) + +let q_unit = CAst.make @@ CTacCst (AbsKn (Tuple 0)) + +let add_generic_scope s entry arg = + let parse = function + | [] -> + let scope = Extend.Aentry entry in + let act x = CAst.make @@ CTacExt (arg, x) in + Tac2entries.ScopeRule (scope, act) + | arg -> scope_fail s arg + in + add_scope s parse + +open CAst + +let () = add_scope "keyword" begin function +| [SexprStr {loc;v=s}] -> + let scope = Extend.Atoken (Tok.PKEYWORD s) in + Tac2entries.ScopeRule (scope, (fun _ -> q_unit)) +| arg -> scope_fail "keyword" arg +end + +let () = add_scope "terminal" begin function +| [SexprStr {loc;v=s}] -> + let scope = Extend.Atoken (CLexer.terminal s) in + Tac2entries.ScopeRule (scope, (fun _ -> q_unit)) +| arg -> scope_fail "terminal" arg +end + +let () = add_scope "list0" begin function +| [tok] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let scope = Extend.Alist0 scope in + let act l = Tac2quote.of_list act l in + Tac2entries.ScopeRule (scope, act) +| [tok; SexprStr {v=str}] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let sep = Extend.Atoken (CLexer.terminal str) in + let scope = Extend.Alist0sep (scope, sep) in + let act l = Tac2quote.of_list act l in + Tac2entries.ScopeRule (scope, act) +| arg -> scope_fail "list0" arg +end + +let () = add_scope "list1" begin function +| [tok] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let scope = Extend.Alist1 scope in + let act l = Tac2quote.of_list act l in + Tac2entries.ScopeRule (scope, act) +| [tok; SexprStr {v=str}] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let sep = Extend.Atoken (CLexer.terminal str) in + let scope = Extend.Alist1sep (scope, sep) in + let act l = Tac2quote.of_list act l in + Tac2entries.ScopeRule (scope, act) +| arg -> scope_fail "list1" arg +end + +let () = add_scope "opt" begin function +| [tok] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let scope = Extend.Aopt scope in + let act opt = match opt with + | None -> + CAst.make @@ CTacCst (AbsKn (Other Core.c_none)) + | Some x -> + CAst.make @@ CTacApp (CAst.make @@ CTacCst (AbsKn (Other Core.c_some)), [act x]) + in + Tac2entries.ScopeRule (scope, act) +| arg -> scope_fail "opt" arg +end + +let () = add_scope "self" begin function +| [] -> + let scope = Extend.Aself in + let act tac = tac in + Tac2entries.ScopeRule (scope, act) +| arg -> scope_fail "self" arg +end + +let () = add_scope "next" begin function +| [] -> + let scope = Extend.Anext in + let act tac = tac in + Tac2entries.ScopeRule (scope, act) +| arg -> scope_fail "next" arg +end + +let () = add_scope "tactic" begin function +| [] -> + (* Default to level 5 parsing *) + let scope = Extend.Aentryl (tac2expr, "5") in + let act tac = tac in + Tac2entries.ScopeRule (scope, act) +| [SexprInt {loc;v=n}] as arg -> + let () = if n < 0 || n > 6 then scope_fail "tactic" arg in + let scope = Extend.Aentryl (tac2expr, string_of_int n) in + let act tac = tac in + Tac2entries.ScopeRule (scope, act) +| arg -> scope_fail "tactic" arg +end + +let () = add_scope "thunk" begin function +| [tok] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let act e = Tac2quote.thunk (act e) in + Tac2entries.ScopeRule (scope, act) +| arg -> scope_fail "thunk" arg +end + +let add_expr_scope name entry f = + add_scope name begin function + | [] -> Tac2entries.ScopeRule (Extend.Aentry entry, f) + | arg -> scope_fail name arg + end + +let () = add_expr_scope "ident" q_ident (fun id -> Tac2quote.of_anti Tac2quote.of_ident id) +let () = add_expr_scope "bindings" q_bindings Tac2quote.of_bindings +let () = add_expr_scope "with_bindings" q_with_bindings Tac2quote.of_bindings +let () = add_expr_scope "intropattern" q_intropattern Tac2quote.of_intro_pattern +let () = add_expr_scope "intropatterns" q_intropatterns Tac2quote.of_intro_patterns +let () = add_expr_scope "destruction_arg" q_destruction_arg Tac2quote.of_destruction_arg +let () = add_expr_scope "induction_clause" q_induction_clause Tac2quote.of_induction_clause +let () = add_expr_scope "conversion" q_conversion Tac2quote.of_conversion +let () = add_expr_scope "rewriting" q_rewriting Tac2quote.of_rewriting +let () = add_expr_scope "clause" q_clause Tac2quote.of_clause +let () = add_expr_scope "hintdb" q_hintdb Tac2quote.of_hintdb +let () = add_expr_scope "occurrences" q_occurrences Tac2quote.of_occurrences +let () = add_expr_scope "dispatch" q_dispatch Tac2quote.of_dispatch +let () = add_expr_scope "strategy" q_strategy_flag Tac2quote.of_strategy_flag +let () = add_expr_scope "reference" q_reference Tac2quote.of_reference +let () = add_expr_scope "move_location" q_move_location Tac2quote.of_move_location +let () = add_expr_scope "pose" q_pose Tac2quote.of_pose +let () = add_expr_scope "assert" q_assert Tac2quote.of_assertion +let () = add_expr_scope "constr_matching" q_constr_matching Tac2quote.of_constr_matching +let () = add_expr_scope "goal_matching" q_goal_matching Tac2quote.of_goal_matching + +let () = add_generic_scope "constr" Pcoq.Constr.constr Tac2quote.wit_constr +let () = add_generic_scope "open_constr" Pcoq.Constr.constr Tac2quote.wit_open_constr +let () = add_generic_scope "pattern" Pcoq.Constr.constr Tac2quote.wit_pattern + +(** seq scope, a bit hairy *) + +open Extend +exception SelfSymbol + +let rec generalize_symbol : + type a tr s. (s, tr, a) Extend.symbol -> (s, Extend.norec, a) Extend.symbol = function +| Atoken tok -> Atoken tok +| Alist1 e -> Alist1 (generalize_symbol e) +| Alist1sep (e, sep) -> + let e = generalize_symbol e in + let sep = generalize_symbol sep in + Alist1sep (e, sep) +| Alist0 e -> Alist0 (generalize_symbol e) +| Alist0sep (e, sep) -> + let e = generalize_symbol e in + let sep = generalize_symbol sep in + Alist0sep (e, sep) +| Aopt e -> Aopt (generalize_symbol e) +| Aself -> raise SelfSymbol +| Anext -> raise SelfSymbol +| Aentry e -> Aentry e +| Aentryl (e, l) -> Aentryl (e, l) +| Arules r -> Arules r + +type _ converter = +| CvNil : (Loc.t -> raw_tacexpr) converter +| CvCns : 'act converter * ('a -> raw_tacexpr) option -> ('a -> 'act) converter + +let rec apply : type a. a converter -> raw_tacexpr list -> a = function +| CvNil -> fun accu loc -> Tac2quote.of_tuple ~loc accu +| CvCns (c, None) -> fun accu x -> apply c accu +| CvCns (c, Some f) -> fun accu x -> apply c (f x :: accu) + +type seqrule = +| Seqrule : (Tac2expr.raw_tacexpr, Extend.norec, 'act, Loc.t -> raw_tacexpr) rule * 'act converter -> seqrule + +let rec make_seq_rule = function +| [] -> + Seqrule (Stop, CvNil) +| tok :: rem -> + let Tac2entries.ScopeRule (scope, f) = Tac2entries.parse_scope tok in + let scope = generalize_symbol scope in + let Seqrule (r, c) = make_seq_rule rem in + let r = NextNoRec (r, scope) in + let f = match tok with + | SexprStr _ -> None (* Leave out mere strings *) + | _ -> Some f + in + Seqrule (r, CvCns (c, f)) + +let () = add_scope "seq" begin fun toks -> + let scope = + try + let Seqrule (r, c) = make_seq_rule (List.rev toks) in + Arules [Rules (r, apply c [])] + with SelfSymbol -> + CErrors.user_err (str "Recursive symbols (self / next) are not allowed in local rules") + in + Tac2entries.ScopeRule (scope, (fun e -> e)) +end diff --git a/vendor/Ltac2/src/tac2core.mli b/vendor/Ltac2/src/tac2core.mli new file mode 100644 index 0000000000..9fae65bb3e --- /dev/null +++ b/vendor/Ltac2/src/tac2core.mli @@ -0,0 +1,30 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Evd.evar_map -> 'a Proofview.tactic) -> 'a Proofview.tactic diff --git a/vendor/Ltac2/src/tac2dyn.ml b/vendor/Ltac2/src/tac2dyn.ml new file mode 100644 index 0000000000..896676f08b --- /dev/null +++ b/vendor/Ltac2/src/tac2dyn.ml @@ -0,0 +1,27 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* ('raw * 'glb) pack + include Arg.Map(struct type 'a t = 'a pack end) +end + +module Val = Dyn.Make(struct end) diff --git a/vendor/Ltac2/src/tac2dyn.mli b/vendor/Ltac2/src/tac2dyn.mli new file mode 100644 index 0000000000..e995296840 --- /dev/null +++ b/vendor/Ltac2/src/tac2dyn.mli @@ -0,0 +1,34 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* ('a, 'b) tag + val eq : ('a1, 'b1) tag -> ('a2, 'b2) tag -> ('a1 * 'b1, 'a2 * 'b2) CSig.eq option + val repr : ('a, 'b) tag -> string +end +(** Arguments that are part of an AST. *) + +module type Param = sig type ('raw, 'glb) t end + +module ArgMap (M : Param) : +sig + type _ pack = Pack : ('raw, 'glb) M.t -> ('raw * 'glb) pack + type t + val empty : t + val add : ('a, 'b) Arg.tag -> ('a * 'b) pack -> t -> t + val remove : ('a, 'b) Arg.tag -> t -> t + val find : ('a, 'b) Arg.tag -> t -> ('a * 'b) pack + val mem : ('a, 'b) Arg.tag -> t -> bool +end + +module Val : Dyn.S +(** Toplevel values *) diff --git a/vendor/Ltac2/src/tac2entries.ml b/vendor/Ltac2/src/tac2entries.ml new file mode 100644 index 0000000000..9fd01426de --- /dev/null +++ b/vendor/Ltac2/src/tac2entries.ml @@ -0,0 +1,938 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* obj = + declare_object {(default_object "TAC2-DEFINITION") with + cache_function = cache_tacdef; + load_function = load_tacdef; + open_function = open_tacdef; + subst_function = subst_tacdef; + classify_function = classify_tacdef} + +(** Type definition *) + +type typdef = { + typdef_local : bool; + typdef_expr : glb_quant_typedef; +} + +let change_kn_label kn id = + let mp = KerName.modpath kn in + KerName.make mp (Label.of_id id) + +let change_sp_label sp id = + let (dp, _) = Libnames.repr_path sp in + Libnames.make_path dp id + +let push_typedef visibility sp kn (_, def) = match def with +| GTydDef _ -> + Tac2env.push_type visibility sp kn +| GTydAlg { galg_constructors = cstrs } -> + (* Register constructors *) + let iter (c, _) = + let spc = change_sp_label sp c in + let knc = change_kn_label kn c in + Tac2env.push_constructor visibility spc knc + in + Tac2env.push_type visibility sp kn; + List.iter iter cstrs +| GTydRec fields -> + (* Register fields *) + let iter (c, _, _) = + let spc = change_sp_label sp c in + let knc = change_kn_label kn c in + Tac2env.push_projection visibility spc knc + in + Tac2env.push_type visibility sp kn; + List.iter iter fields +| GTydOpn -> + Tac2env.push_type visibility sp kn + +let next i = + let ans = !i in + let () = incr i in + ans + +let define_typedef kn (params, def as qdef) = match def with +| GTydDef _ -> + Tac2env.define_type kn qdef +| GTydAlg { galg_constructors = cstrs } -> + (* Define constructors *) + let constant = ref 0 in + let nonconstant = ref 0 in + let iter (c, args) = + let knc = change_kn_label kn c in + let tag = if List.is_empty args then next constant else next nonconstant in + let data = { + Tac2env.cdata_prms = params; + cdata_type = kn; + cdata_args = args; + cdata_indx = Some tag; + } in + Tac2env.define_constructor knc data + in + Tac2env.define_type kn qdef; + List.iter iter cstrs +| GTydRec fs -> + (* Define projections *) + let iter i (id, mut, t) = + let knp = change_kn_label kn id in + let proj = { + Tac2env.pdata_prms = params; + pdata_type = kn; + pdata_ptyp = t; + pdata_mutb = mut; + pdata_indx = i; + } in + Tac2env.define_projection knp proj + in + Tac2env.define_type kn qdef; + List.iteri iter fs +| GTydOpn -> + Tac2env.define_type kn qdef + +let perform_typdef vs ((sp, kn), def) = + let () = if not def.typdef_local then push_typedef vs sp kn def.typdef_expr in + define_typedef kn def.typdef_expr + +let load_typdef i obj = perform_typdef (Until i) obj +let open_typdef i obj = perform_typdef (Exactly i) obj + +let cache_typdef ((sp, kn), def) = + let () = push_typedef (Until 1) sp kn def.typdef_expr in + define_typedef kn def.typdef_expr + +let subst_typdef (subst, def) = + let expr' = subst_quant_typedef subst def.typdef_expr in + if expr' == def.typdef_expr then def else { def with typdef_expr = expr' } + +let classify_typdef o = Substitute o + +let inTypDef : typdef -> obj = + declare_object {(default_object "TAC2-TYPE-DEFINITION") with + cache_function = cache_typdef; + load_function = load_typdef; + open_function = open_typdef; + subst_function = subst_typdef; + classify_function = classify_typdef} + +(** Type extension *) + +type extension_data = { + edata_name : Id.t; + edata_args : int glb_typexpr list; +} + +type typext = { + typext_local : bool; + typext_prms : int; + typext_type : type_constant; + typext_expr : extension_data list; +} + +let push_typext vis sp kn def = + let iter data = + let spc = change_sp_label sp data.edata_name in + let knc = change_kn_label kn data.edata_name in + Tac2env.push_constructor vis spc knc + in + List.iter iter def.typext_expr + +let define_typext kn def = + let iter data = + let knc = change_kn_label kn data.edata_name in + let cdata = { + Tac2env.cdata_prms = def.typext_prms; + cdata_type = def.typext_type; + cdata_args = data.edata_args; + cdata_indx = None; + } in + Tac2env.define_constructor knc cdata + in + List.iter iter def.typext_expr + +let cache_typext ((sp, kn), def) = + let () = define_typext kn def in + push_typext (Until 1) sp kn def + +let perform_typext vs ((sp, kn), def) = + let () = if not def.typext_local then push_typext vs sp kn def in + define_typext kn def + +let load_typext i obj = perform_typext (Until i) obj +let open_typext i obj = perform_typext (Exactly i) obj + +let subst_typext (subst, e) = + let open Mod_subst in + let subst_data data = + let edata_args = List.Smart.map (fun e -> subst_type subst e) data.edata_args in + if edata_args == data.edata_args then data + else { data with edata_args } + in + let typext_type = subst_kn subst e.typext_type in + let typext_expr = List.Smart.map subst_data e.typext_expr in + if typext_type == e.typext_type && typext_expr == e.typext_expr then + e + else + { e with typext_type; typext_expr } + +let classify_typext o = Substitute o + +let inTypExt : typext -> obj = + declare_object {(default_object "TAC2-TYPE-EXTENSION") with + cache_function = cache_typext; + load_function = load_typext; + open_function = open_typext; + subst_function = subst_typext; + classify_function = classify_typext} + +(** Toplevel entries *) + +let fresh_var avoid x = + let bad id = + Id.Set.mem id avoid || + (try ignore (Tac2env.locate_ltac (qualid_of_ident id)); true with Not_found -> false) + in + Namegen.next_ident_away_from (Id.of_string x) bad + +let extract_pattern_type ({loc;v=p} as pat) = match p with +| CPatCnv (pat, ty) -> pat, Some ty +| CPatVar _ | CPatRef _ -> pat, None + +(** Mangle recursive tactics *) +let inline_rec_tactic tactics = + let avoid = List.fold_left (fun accu ({v=id}, _) -> Id.Set.add id accu) Id.Set.empty tactics in + let map (id, e) = match e.v with + | CTacFun (pat, _) -> (id, List.map extract_pattern_type pat, e) + | _ -> + user_err ?loc:id.loc (str "Recursive tactic definitions must be functions") + in + let tactics = List.map map tactics in + let map (id, pat, e) = + let fold_var (avoid, ans) (pat, _) = + let id = fresh_var avoid "x" in + let loc = pat.loc in + (Id.Set.add id avoid, CAst.make ?loc id :: ans) + in + (* Fresh variables to abstract over the function patterns *) + let _, vars = List.fold_left fold_var (avoid, []) pat in + let map_body ({loc;v=id}, _, e) = CAst.(make ?loc @@ CPatVar (Name id)), e in + let bnd = List.map map_body tactics in + let pat_of_id {loc;v=id} = CAst.make ?loc @@ CPatVar (Name id) in + let var_of_id {loc;v=id} = + let qid = qualid_of_ident ?loc id in + CAst.make ?loc @@ CTacRef (RelId qid) + in + let loc0 = e.loc in + let vpat = List.map pat_of_id vars in + let varg = List.map var_of_id vars in + let e = CAst.make ?loc:loc0 @@ CTacLet (true, bnd, CAst.make ?loc:loc0 @@ CTacApp (var_of_id id, varg)) in + (id, CAst.make ?loc:loc0 @@ CTacFun (vpat, e)) + in + List.map map tactics + +let check_lowercase {loc;v=id} = + if Tac2env.is_constructor (Libnames.qualid_of_ident id) then + user_err ?loc (str "The identifier " ++ Id.print id ++ str " must be lowercase") + +let register_ltac ?(local = false) ?(mut = false) isrec tactics = + let map ({loc;v=na}, e) = + let id = match na with + | Anonymous -> + user_err ?loc (str "Tactic definition must have a name") + | Name id -> id + in + let () = check_lowercase CAst.(make ?loc id) in + (CAst.(make ?loc id), e) + in + let tactics = List.map map tactics in + let tactics = + if isrec then inline_rec_tactic tactics else tactics + in + let map ({loc;v=id}, e) = + let (e, t) = intern ~strict:true e in + let () = + if not (is_value e) then + user_err ?loc (str "Tactic definition must be a syntactical value") + in + let kn = Lib.make_kn id in + let exists = + try let _ = Tac2env.interp_global kn in true with Not_found -> false + in + let () = + if exists then + user_err ?loc (str "Tactic " ++ Names.Id.print id ++ str " already exists") + in + (id, e, t) + in + let defs = List.map map tactics in + let iter (id, e, t) = + let def = { + tacdef_local = local; + tacdef_mutable = mut; + tacdef_expr = e; + tacdef_type = t; + } in + ignore (Lib.add_leaf id (inTacDef def)) + in + List.iter iter defs + +let qualid_to_ident qid = + if qualid_is_ident qid then CAst.make ?loc:qid.CAst.loc @@ qualid_basename qid + else user_err ?loc:qid.CAst.loc (str "Identifier expected") + +let register_typedef ?(local = false) isrec types = + let same_name ({v=id1}, _) ({v=id2}, _) = Id.equal id1 id2 in + let () = match List.duplicates same_name types with + | [] -> () + | ({loc;v=id}, _) :: _ -> + user_err ?loc (str "Multiple definition of the type name " ++ Id.print id) + in + let check ({loc;v=id}, (params, def)) = + let same_name {v=id1} {v=id2} = Id.equal id1 id2 in + let () = match List.duplicates same_name params with + | [] -> () + | {loc;v=id} :: _ -> + user_err ?loc (str "The type parameter " ++ Id.print id ++ + str " occurs several times") + in + match def with + | CTydDef _ -> + if isrec then + user_err ?loc (str "The type abbreviation " ++ Id.print id ++ + str " cannot be recursive") + | CTydAlg cs -> + let same_name (id1, _) (id2, _) = Id.equal id1 id2 in + let () = match List.duplicates same_name cs with + | [] -> () + | (id, _) :: _ -> + user_err (str "Multiple definitions of the constructor " ++ Id.print id) + in + () + | CTydRec ps -> + let same_name (id1, _, _) (id2, _, _) = Id.equal id1 id2 in + let () = match List.duplicates same_name ps with + | [] -> () + | (id, _, _) :: _ -> + user_err (str "Multiple definitions of the projection " ++ Id.print id) + in + () + | CTydOpn -> + if isrec then + user_err ?loc (str "The open type declaration " ++ Id.print id ++ + str " cannot be recursive") + in + let () = List.iter check types in + let self = + if isrec then + let fold accu ({v=id}, (params, _)) = + Id.Map.add id (Lib.make_kn id, List.length params) accu + in + List.fold_left fold Id.Map.empty types + else Id.Map.empty + in + let map ({v=id}, def) = + let typdef = { + typdef_local = local; + typdef_expr = intern_typedef self def; + } in + (id, typdef) + in + let types = List.map map types in + let iter (id, def) = ignore (Lib.add_leaf id (inTypDef def)) in + List.iter iter types + +let register_primitive ?(local = false) {loc;v=id} t ml = + let t = intern_open_type t in + let rec count_arrow = function + | GTypArrow (_, t) -> 1 + count_arrow t + | _ -> 0 + in + let arrows = count_arrow (snd t) in + let () = if Int.equal arrows 0 then + user_err ?loc (str "External tactic must have at least one argument") in + let () = + try let _ = Tac2env.interp_primitive ml in () with Not_found -> + user_err ?loc (str "Unregistered primitive " ++ + quote (str ml.mltac_plugin) ++ spc () ++ quote (str ml.mltac_tactic)) + in + let init i = Id.of_string (Printf.sprintf "x%i" i) in + let names = List.init arrows init in + let bnd = List.map (fun id -> Name id) names in + let arg = List.map (fun id -> GTacVar id) names in + let e = GTacFun (bnd, GTacPrm (ml, arg)) in + let def = { + tacdef_local = local; + tacdef_mutable = false; + tacdef_expr = e; + tacdef_type = t; + } in + ignore (Lib.add_leaf id (inTacDef def)) + +let register_open ?(local = false) qid (params, def) = + let kn = + try Tac2env.locate_type qid + with Not_found -> + user_err ?loc:qid.CAst.loc (str "Unbound type " ++ pr_qualid qid) + in + let (tparams, t) = Tac2env.interp_type kn in + let () = match t with + | GTydOpn -> () + | GTydAlg _ | GTydRec _ | GTydDef _ -> + user_err ?loc:qid.CAst.loc (str "Type " ++ pr_qualid qid ++ str " is not an open type") + in + let () = + if not (Int.equal (List.length params) tparams) then + Tac2intern.error_nparams_mismatch ?loc:qid.CAst.loc (List.length params) tparams + in + match def with + | CTydOpn -> () + | CTydAlg def -> + let intern_type t = + let tpe = CTydDef (Some t) in + let (_, ans) = intern_typedef Id.Map.empty (params, tpe) in + match ans with + | GTydDef (Some t) -> t + | _ -> assert false + in + let map (id, tpe) = + let tpe = List.map intern_type tpe in + { edata_name = id; edata_args = tpe } + in + let def = List.map map def in + let def = { + typext_local = local; + typext_type = kn; + typext_prms = tparams; + typext_expr = def; + } in + Lib.add_anonymous_leaf (inTypExt def) + | CTydRec _ | CTydDef _ -> + user_err ?loc:qid.CAst.loc (str "Extensions only accept inductive constructors") + +let register_type ?local isrec types = match types with +| [qid, true, def] -> + let () = if isrec then user_err ?loc:qid.CAst.loc (str "Extensions cannot be recursive") in + register_open ?local qid def +| _ -> + let map (qid, redef, def) = + let () = if redef then + user_err ?loc:qid.loc (str "Types can only be extended one by one") + in + (qualid_to_ident qid, def) + in + let types = List.map map types in + register_typedef ?local isrec types + +(** Parsing *) + +type 'a token = +| TacTerm of string +| TacNonTerm of Name.t * 'a + +type scope_rule = +| ScopeRule : (raw_tacexpr, _, 'a) Extend.symbol * ('a -> raw_tacexpr) -> scope_rule + +type scope_interpretation = sexpr list -> scope_rule + +let scope_table : scope_interpretation Id.Map.t ref = ref Id.Map.empty + +let register_scope id s = + scope_table := Id.Map.add id s !scope_table + +module ParseToken = +struct + +let loc_of_token = function +| SexprStr {loc} -> loc +| SexprInt {loc} -> loc +| SexprRec (loc, _, _) -> Some loc + +let parse_scope = function +| SexprRec (_, {loc;v=Some id}, toks) -> + if Id.Map.mem id !scope_table then + Id.Map.find id !scope_table toks + else + CErrors.user_err ?loc (str "Unknown scope" ++ spc () ++ Names.Id.print id) +| SexprStr {v=str} -> + let v_unit = CAst.make @@ CTacCst (AbsKn (Tuple 0)) in + ScopeRule (Extend.Atoken (Tok.PIDENT (Some str)), (fun _ -> v_unit)) +| tok -> + let loc = loc_of_token tok in + CErrors.user_err ?loc (str "Invalid parsing token") + +let parse_token = function +| SexprStr {v=s} -> TacTerm s +| SexprRec (_, {v=na}, [tok]) -> + let na = match na with None -> Anonymous | Some id -> Name id in + let scope = parse_scope tok in + TacNonTerm (na, scope) +| tok -> + let loc = loc_of_token tok in + CErrors.user_err ?loc (str "Invalid parsing token") + +end + +let parse_scope = ParseToken.parse_scope + +type synext = { + synext_tok : sexpr list; + synext_exp : raw_tacexpr; + synext_lev : int option; + synext_loc : bool; +} + +type krule = +| KRule : + (raw_tacexpr, _, 'act, Loc.t -> raw_tacexpr) Extend.rule * + ((Loc.t -> (Name.t * raw_tacexpr) list -> raw_tacexpr) -> 'act) -> krule + +let rec get_rule (tok : scope_rule token list) : krule = match tok with +| [] -> KRule (Extend.Stop, fun k loc -> k loc []) +| TacNonTerm (na, ScopeRule (scope, inj)) :: tok -> + let KRule (rule, act) = get_rule tok in + let rule = Extend.Next (rule, scope) in + let act k e = act (fun loc acc -> k loc ((na, inj e) :: acc)) in + KRule (rule, act) +| TacTerm t :: tok -> + let KRule (rule, act) = get_rule tok in + let rule = Extend.Next (rule, Extend.Atoken (CLexer.terminal t)) in + let act k _ = act k in + KRule (rule, act) + +let perform_notation syn st = + let tok = List.rev_map ParseToken.parse_token syn.synext_tok in + let KRule (rule, act) = get_rule tok in + let mk loc args = + let map (na, e) = + ((CAst.make ?loc:e.loc @@ CPatVar na), e) + in + let bnd = List.map map args in + CAst.make ~loc @@ CTacLet (false, bnd, syn.synext_exp) + in + let rule = Extend.Rule (rule, act mk) in + let lev = match syn.synext_lev with + | None -> None + | Some lev -> Some (string_of_int lev) + in + let rule = (lev, None, [rule]) in + ([Pcoq.ExtendRule (Pltac.tac2expr, None, (None, [rule]))], st) + +let ltac2_notation = + Pcoq.create_grammar_command "ltac2-notation" perform_notation + +let cache_synext (_, syn) = + Pcoq.extend_grammar_command ltac2_notation syn + +let open_synext i (_, syn) = + if Int.equal i 1 then Pcoq.extend_grammar_command ltac2_notation syn + +let subst_synext (subst, syn) = + let e = Tac2intern.subst_rawexpr subst syn.synext_exp in + if e == syn.synext_exp then syn else { syn with synext_exp = e } + +let classify_synext o = + if o.synext_loc then Dispose else Substitute o + +let inTac2Notation : synext -> obj = + declare_object {(default_object "TAC2-NOTATION") with + cache_function = cache_synext; + open_function = open_synext; + subst_function = subst_synext; + classify_function = classify_synext} + +type abbreviation = { + abbr_body : raw_tacexpr; +} + +let perform_abbreviation visibility ((sp, kn), abbr) = + let () = Tac2env.push_ltac visibility sp (TacAlias kn) in + Tac2env.define_alias kn abbr.abbr_body + +let load_abbreviation i obj = perform_abbreviation (Until i) obj +let open_abbreviation i obj = perform_abbreviation (Exactly i) obj + +let cache_abbreviation ((sp, kn), abbr) = + let () = Tac2env.push_ltac (Until 1) sp (TacAlias kn) in + Tac2env.define_alias kn abbr.abbr_body + +let subst_abbreviation (subst, abbr) = + let body' = subst_rawexpr subst abbr.abbr_body in + if body' == abbr.abbr_body then abbr + else { abbr_body = body' } + +let classify_abbreviation o = Substitute o + +let inTac2Abbreviation : abbreviation -> obj = + declare_object {(default_object "TAC2-ABBREVIATION") with + cache_function = cache_abbreviation; + load_function = load_abbreviation; + open_function = open_abbreviation; + subst_function = subst_abbreviation; + classify_function = classify_abbreviation} + +let register_notation ?(local = false) tkn lev body = match tkn, lev with +| [SexprRec (_, {loc;v=Some id}, [])], None -> + (* Tactic abbreviation *) + let () = check_lowercase CAst.(make ?loc id) in + let body = Tac2intern.globalize Id.Set.empty body in + let abbr = { abbr_body = body } in + ignore (Lib.add_leaf id (inTac2Abbreviation abbr)) +| _ -> + (* Check that the tokens make sense *) + let entries = List.map ParseToken.parse_token tkn in + let fold accu tok = match tok with + | TacTerm _ -> accu + | TacNonTerm (Name id, _) -> Id.Set.add id accu + | TacNonTerm (Anonymous, _) -> accu + in + let ids = List.fold_left fold Id.Set.empty entries in + (* Globalize so that names are absolute *) + let body = Tac2intern.globalize ids body in + let lev = match lev with Some _ -> lev | None -> Some 5 in + let ext = { + synext_tok = tkn; + synext_exp = body; + synext_lev = lev; + synext_loc = local; + } in + Lib.add_anonymous_leaf (inTac2Notation ext) + +type redefinition = { + redef_kn : ltac_constant; + redef_body : glb_tacexpr; +} + +let perform_redefinition (_, redef) = + let kn = redef.redef_kn in + let data = Tac2env.interp_global kn in + let data = { data with Tac2env.gdata_expr = redef.redef_body } in + Tac2env.define_global kn data + +let subst_redefinition (subst, redef) = + let kn = Mod_subst.subst_kn subst redef.redef_kn in + let body = Tac2intern.subst_expr subst redef.redef_body in + if kn == redef.redef_kn && body == redef.redef_body then redef + else { redef_kn = kn; redef_body = body } + +let classify_redefinition o = Substitute o + +let inTac2Redefinition : redefinition -> obj = + declare_object {(default_object "TAC2-REDEFINITION") with + cache_function = perform_redefinition; + open_function = (fun _ -> perform_redefinition); + subst_function = subst_redefinition; + classify_function = classify_redefinition } + +let register_redefinition ?(local = false) qid e = + let kn = + try Tac2env.locate_ltac qid + with Not_found -> user_err ?loc:qid.CAst.loc (str "Unknown tactic " ++ pr_qualid qid) + in + let kn = match kn with + | TacConstant kn -> kn + | TacAlias _ -> + user_err ?loc:qid.CAst.loc (str "Cannot redefine syntactic abbreviations") + in + let data = Tac2env.interp_global kn in + let () = + if not (data.Tac2env.gdata_mutable) then + user_err ?loc:qid.CAst.loc (str "The tactic " ++ pr_qualid qid ++ str " is not declared as mutable") + in + let (e, t) = intern ~strict:true e in + let () = + if not (is_value e) then + user_err ?loc:qid.CAst.loc (str "Tactic definition must be a syntactical value") + in + let () = + if not (Tac2intern.check_subtype t data.Tac2env.gdata_type) then + let name = int_name () in + user_err ?loc:qid.CAst.loc (str "Type " ++ pr_glbtype name (snd t) ++ + str " is not a subtype of " ++ pr_glbtype name (snd data.Tac2env.gdata_type)) + in + let def = { + redef_kn = kn; + redef_body = e; + } in + Lib.add_anonymous_leaf (inTac2Redefinition def) + +let perform_eval ~pstate e = + let open Proofview.Notations in + let env = Global.env () in + let (e, ty) = Tac2intern.intern ~strict:false e in + let v = Tac2interp.interp Tac2interp.empty_environment e in + let selector, proof = + match pstate with + | None -> + let sigma = Evd.from_env env in + let name, poly = Id.of_string "ltac2", false in + Goal_select.SelectAll, Proof.start ~name ~poly sigma [] + | Some pstate -> + Goal_select.get_default_goal_selector (), + Proof_global.give_me_the_proof pstate + in + let v = match selector with + | Goal_select.SelectNth i -> Proofview.tclFOCUS i i v + | Goal_select.SelectList l -> Proofview.tclFOCUSLIST l v + | Goal_select.SelectId id -> Proofview.tclFOCUSID id v + | Goal_select.SelectAll -> v + | Goal_select.SelectAlreadyFocused -> assert false (* TODO **) + in + (* HACK: the API doesn't allow to return a value *) + let ans = ref None in + let tac = (v >>= fun r -> ans := Some r; Proofview.tclUNIT ()) in + let (proof, _) = Proof.run_tactic (Global.env ()) tac proof in + let sigma = Proof.in_proof proof (fun sigma -> sigma) in + let ans = match !ans with None -> assert false | Some r -> r in + let name = int_name () in + Feedback.msg_notice (str "- : " ++ pr_glbtype name (snd ty) + ++ spc () ++ str "=" ++ spc () ++ + Tac2print.pr_valexpr env sigma ans (snd ty)) + +(** Toplevel entries *) + +let register_struct ?local ~pstate str = match str with +| StrVal (mut, isrec, e) -> register_ltac ?local ~mut isrec e +| StrTyp (isrec, t) -> register_type ?local isrec t +| StrPrm (id, t, ml) -> register_primitive ?local id t ml +| StrSyn (tok, lev, e) -> register_notation ?local tok lev e +| StrMut (qid, e) -> register_redefinition ?local qid e +| StrRun e -> perform_eval ~pstate e + +(** Toplevel exception *) + +let _ = Goptions.declare_bool_option { + Goptions.optdepr = false; + Goptions.optname = "print Ltac2 backtrace"; + Goptions.optkey = ["Ltac2"; "Backtrace"]; + Goptions.optread = (fun () -> !Tac2interp.print_ltac2_backtrace); + Goptions.optwrite = (fun b -> Tac2interp.print_ltac2_backtrace := b); +} + +let backtrace : backtrace Exninfo.t = Exninfo.make () + +let pr_frame = function +| FrAnon e -> str "Call {" ++ pr_glbexpr e ++ str "}" +| FrLtac kn -> + str "Call " ++ Libnames.pr_qualid (Tac2env.shortest_qualid_of_ltac (TacConstant kn)) +| FrPrim ml -> + str "Prim <" ++ str ml.mltac_plugin ++ str ":" ++ str ml.mltac_tactic ++ str ">" +| FrExtn (tag, arg) -> + let obj = Tac2env.interp_ml_object tag in + str "Extn " ++ str (Tac2dyn.Arg.repr tag) ++ str ":" ++ spc () ++ + obj.Tac2env.ml_print (Global.env ()) arg + +let () = register_handler begin function +| Tac2interp.LtacError (kn, args) -> + let t_exn = KerName.make Tac2env.coq_prefix (Label.make "exn") in + let v = Tac2ffi.of_open (kn, args) in + let t = GTypRef (Other t_exn, []) in + let c = Tac2print.pr_valexpr (Global.env ()) Evd.empty v t in + hov 0 (str "Uncaught Ltac2 exception:" ++ spc () ++ hov 0 c) +| _ -> raise Unhandled +end + +let () = ExplainErr.register_additional_error_info begin fun (e, info) -> + if !Tac2interp.print_ltac2_backtrace then + let bt = Exninfo.get info backtrace in + let bt = match bt with + | Some bt -> bt + | None -> raise Exit + in + let bt = + str "Backtrace:" ++ fnl () ++ prlist_with_sep fnl pr_frame bt ++ fnl () + in + Some (Loc.tag @@ Some bt) + else raise Exit +end + +(** Printing *) + +let print_ltac qid = + if Tac2env.is_constructor qid then + let kn = + try Tac2env.locate_constructor qid + with Not_found -> user_err ?loc:qid.CAst.loc (str "Unknown constructor " ++ pr_qualid qid) + in + let _ = Tac2env.interp_constructor kn in + Feedback.msg_notice (hov 2 (str "Constructor" ++ spc () ++ str ":" ++ spc () ++ pr_qualid qid)) + else + let kn = + try Tac2env.locate_ltac qid + with Not_found -> user_err ?loc:qid.CAst.loc (str "Unknown tactic " ++ pr_qualid qid) + in + match kn with + | TacConstant kn -> + let data = Tac2env.interp_global kn in + let e = data.Tac2env.gdata_expr in + let (_, t) = data.Tac2env.gdata_type in + let name = int_name () in + Feedback.msg_notice ( + hov 0 ( + hov 2 (pr_qualid qid ++ spc () ++ str ":" ++ spc () ++ pr_glbtype name t) ++ fnl () ++ + hov 2 (pr_qualid qid ++ spc () ++ str ":=" ++ spc () ++ pr_glbexpr e) + ) + ) + | TacAlias kn -> + Feedback.msg_notice (str "Alias to ...") + +(** Calling tactics *) + +let solve ~pstate default tac = + let pstate, status = Proof_global.with_current_proof begin fun etac p -> + let with_end_tac = if default then Some etac else None in + let g = Goal_select.get_default_goal_selector () in + let (p, status) = Pfedit.solve g None tac ?with_end_tac p in + (* in case a strict subtree was completed, + go back to the top of the prooftree *) + let p = Proof.maximal_unfocus Vernacentries.command_focus p in + p, status + end pstate in + if not status then Feedback.feedback Feedback.AddedAxiom; + pstate + +let call ~pstate ~default e = + let loc = e.loc in + let (e, t) = intern ~strict:false e in + let () = check_unit ?loc t in + let tac = Tac2interp.interp Tac2interp.empty_environment e in + solve ~pstate default (Proofview.tclIGNORE tac) + +(** Primitive algebraic types than can't be defined Coq-side *) + +let register_prim_alg name params def = + let id = Id.of_string name in + let def = List.map (fun (cstr, tpe) -> (Id.of_string_soft cstr, tpe)) def in + let getn (const, nonconst) (c, args) = match args with + | [] -> (succ const, nonconst) + | _ :: _ -> (const, succ nonconst) + in + let nconst, nnonconst = List.fold_left getn (0, 0) def in + let alg = { + galg_constructors = def; + galg_nconst = nconst; + galg_nnonconst = nnonconst; + } in + let def = (params, GTydAlg alg) in + let def = { typdef_local = false; typdef_expr = def } in + ignore (Lib.add_leaf id (inTypDef def)) + +let coq_def n = KerName.make Tac2env.coq_prefix (Label.make n) + +let def_unit = { + typdef_local = false; + typdef_expr = 0, GTydDef (Some (GTypRef (Tuple 0, []))); +} + +let t_list = coq_def "list" + +let (f_register_constr_quotations, register_constr_quotations) = Hook.make () + +let cache_ltac2_init (_, ()) = + Hook.get f_register_constr_quotations () + +let load_ltac2_init _ (_, ()) = + Hook.get f_register_constr_quotations () + +let open_ltac2_init _ (_, ()) = + Goptions.set_string_option_value_gen ["Default"; "Proof"; "Mode"] "Ltac2" + +(** Dummy object that register global rules when Require is called *) +let inTac2Init : unit -> obj = + declare_object {(default_object "TAC2-INIT") with + cache_function = cache_ltac2_init; + load_function = load_ltac2_init; + open_function = open_ltac2_init; + } + +let _ = Mltop.declare_cache_obj begin fun () -> + ignore (Lib.add_leaf (Id.of_string "unit") (inTypDef def_unit)); + register_prim_alg "list" 1 [ + ("[]", []); + ("::", [GTypVar 0; GTypRef (Other t_list, [GTypVar 0])]); + ]; + Lib.add_anonymous_leaf (inTac2Init ()); +end "ltac2_plugin" diff --git a/vendor/Ltac2/src/tac2entries.mli b/vendor/Ltac2/src/tac2entries.mli new file mode 100644 index 0000000000..d493192bb3 --- /dev/null +++ b/vendor/Ltac2/src/tac2entries.mli @@ -0,0 +1,93 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* ?mut:bool -> rec_flag -> + (Names.lname * raw_tacexpr) list -> unit + +val register_type : ?local:bool -> rec_flag -> + (qualid * redef_flag * raw_quant_typedef) list -> unit + +val register_primitive : ?local:bool -> + Names.lident -> raw_typexpr -> ml_tactic_name -> unit + +val register_struct + : ?local:bool + -> pstate:Proof_global.t option + -> strexpr + -> unit + +val register_notation : ?local:bool -> sexpr list -> int option -> + raw_tacexpr -> unit + +(** {5 Notations} *) + +type scope_rule = +| ScopeRule : (raw_tacexpr, _, 'a) Extend.symbol * ('a -> raw_tacexpr) -> scope_rule + +type scope_interpretation = sexpr list -> scope_rule + +val register_scope : Id.t -> scope_interpretation -> unit +(** Create a new scope with the provided name *) + +val parse_scope : sexpr -> scope_rule +(** Use this to interpret the subscopes for interpretation functions *) + +(** {5 Inspecting} *) + +val print_ltac : Libnames.qualid -> unit + +(** {5 Eval loop} *) + +(** Evaluate a tactic expression in the current environment *) +val call : pstate:Proof_global.t -> default:bool -> raw_tacexpr -> Proof_global.t + +(** {5 Toplevel exceptions} *) + +val backtrace : backtrace Exninfo.t + +(** {5 Parsing entries} *) + +module Pltac : +sig +val tac2expr : raw_tacexpr Pcoq.Entry.t + +(** Quoted entries. To be used for complex notations. *) + +open Tac2qexpr + +val q_ident : Id.t CAst.t or_anti Pcoq.Entry.t +val q_bindings : bindings Pcoq.Entry.t +val q_with_bindings : bindings Pcoq.Entry.t +val q_intropattern : intro_pattern Pcoq.Entry.t +val q_intropatterns : intro_pattern list CAst.t Pcoq.Entry.t +val q_destruction_arg : destruction_arg Pcoq.Entry.t +val q_induction_clause : induction_clause Pcoq.Entry.t +val q_conversion : conversion Pcoq.Entry.t +val q_rewriting : rewriting Pcoq.Entry.t +val q_clause : clause Pcoq.Entry.t +val q_dispatch : dispatch Pcoq.Entry.t +val q_occurrences : occurrences Pcoq.Entry.t +val q_reference : reference or_anti Pcoq.Entry.t +val q_strategy_flag : strategy_flag Pcoq.Entry.t +val q_constr_matching : constr_matching Pcoq.Entry.t +val q_goal_matching : goal_matching Pcoq.Entry.t +val q_hintdb : hintdb Pcoq.Entry.t +val q_move_location : move_location Pcoq.Entry.t +val q_pose : pose Pcoq.Entry.t +val q_assert : assertion Pcoq.Entry.t +end + +(** {5 Hooks} *) + +val register_constr_quotations : (unit -> unit) Hook.t diff --git a/vendor/Ltac2/src/tac2env.ml b/vendor/Ltac2/src/tac2env.ml new file mode 100644 index 0000000000..93ad57e97e --- /dev/null +++ b/vendor/Ltac2/src/tac2env.ml @@ -0,0 +1,298 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* KerName.compare c1 c2 +| TacAlias c1, TacAlias c2 -> KerName.compare c1 c2 +| TacConstant _, TacAlias _ -> -1 +| TacAlias _, TacConstant _ -> 1 + +let equal r1 r2 = compare r1 r2 == 0 + +end + +module KnTab = Nametab.Make(FullPath)(KerName) +module RfTab = Nametab.Make(FullPath)(TacRef) +module RfMap = Map.Make(TacRef) + +type nametab = { + tab_ltac : RfTab.t; + tab_ltac_rev : full_path RfMap.t; + tab_cstr : KnTab.t; + tab_cstr_rev : full_path KNmap.t; + tab_type : KnTab.t; + tab_type_rev : full_path KNmap.t; + tab_proj : KnTab.t; + tab_proj_rev : full_path KNmap.t; +} + +let empty_nametab = { + tab_ltac = RfTab.empty; + tab_ltac_rev = RfMap.empty; + tab_cstr = KnTab.empty; + tab_cstr_rev = KNmap.empty; + tab_type = KnTab.empty; + tab_type_rev = KNmap.empty; + tab_proj = KnTab.empty; + tab_proj_rev = KNmap.empty; +} + +let nametab = Summary.ref empty_nametab ~name:"ltac2-nametab" + +let push_ltac vis sp kn = + let tab = !nametab in + let tab_ltac = RfTab.push vis sp kn tab.tab_ltac in + let tab_ltac_rev = RfMap.add kn sp tab.tab_ltac_rev in + nametab := { tab with tab_ltac; tab_ltac_rev } + +let locate_ltac qid = + let tab = !nametab in + RfTab.locate qid tab.tab_ltac + +let locate_extended_all_ltac qid = + let tab = !nametab in + RfTab.find_prefixes qid tab.tab_ltac + +let shortest_qualid_of_ltac kn = + let tab = !nametab in + let sp = RfMap.find kn tab.tab_ltac_rev in + RfTab.shortest_qualid Id.Set.empty sp tab.tab_ltac + +let push_constructor vis sp kn = + let tab = !nametab in + let tab_cstr = KnTab.push vis sp kn tab.tab_cstr in + let tab_cstr_rev = KNmap.add kn sp tab.tab_cstr_rev in + nametab := { tab with tab_cstr; tab_cstr_rev } + +let locate_constructor qid = + let tab = !nametab in + KnTab.locate qid tab.tab_cstr + +let locate_extended_all_constructor qid = + let tab = !nametab in + KnTab.find_prefixes qid tab.tab_cstr + +let shortest_qualid_of_constructor kn = + let tab = !nametab in + let sp = KNmap.find kn tab.tab_cstr_rev in + KnTab.shortest_qualid Id.Set.empty sp tab.tab_cstr + +let push_type vis sp kn = + let tab = !nametab in + let tab_type = KnTab.push vis sp kn tab.tab_type in + let tab_type_rev = KNmap.add kn sp tab.tab_type_rev in + nametab := { tab with tab_type; tab_type_rev } + +let locate_type qid = + let tab = !nametab in + KnTab.locate qid tab.tab_type + +let locate_extended_all_type qid = + let tab = !nametab in + KnTab.find_prefixes qid tab.tab_type + +let shortest_qualid_of_type ?loc kn = + let tab = !nametab in + let sp = KNmap.find kn tab.tab_type_rev in + KnTab.shortest_qualid ?loc Id.Set.empty sp tab.tab_type + +let push_projection vis sp kn = + let tab = !nametab in + let tab_proj = KnTab.push vis sp kn tab.tab_proj in + let tab_proj_rev = KNmap.add kn sp tab.tab_proj_rev in + nametab := { tab with tab_proj; tab_proj_rev } + +let locate_projection qid = + let tab = !nametab in + KnTab.locate qid tab.tab_proj + +let locate_extended_all_projection qid = + let tab = !nametab in + KnTab.find_prefixes qid tab.tab_proj + +let shortest_qualid_of_projection kn = + let tab = !nametab in + let sp = KNmap.find kn tab.tab_proj_rev in + KnTab.shortest_qualid Id.Set.empty sp tab.tab_proj + +type 'a or_glb_tacexpr = +| GlbVal of 'a +| GlbTacexpr of glb_tacexpr + +type environment = { + env_ist : valexpr Id.Map.t; +} + +type ('a, 'b, 'r) intern_fun = Genintern.glob_sign -> 'a -> 'b * 'r glb_typexpr + +type ('a, 'b) ml_object = { + ml_intern : 'r. (raw_tacexpr, glb_tacexpr, 'r) intern_fun -> ('a, 'b or_glb_tacexpr, 'r) intern_fun; + ml_subst : Mod_subst.substitution -> 'b -> 'b; + ml_interp : environment -> 'b -> valexpr Proofview.tactic; + ml_print : Environ.env -> 'b -> Pp.t; +} + +module MLTypeObj = +struct + type ('a, 'b) t = ('a, 'b) ml_object +end + +module MLType = Tac2dyn.ArgMap(MLTypeObj) + +let ml_object_table = ref MLType.empty + +let define_ml_object t tpe = + ml_object_table := MLType.add t (MLType.Pack tpe) !ml_object_table + +let interp_ml_object t = + try + let MLType.Pack ans = MLType.find t !ml_object_table in + ans + with Not_found -> + CErrors.anomaly Pp.(str "Unknown object type " ++ str (Tac2dyn.Arg.repr t)) + +(** Absolute paths *) + +let coq_prefix = + MPfile (DirPath.make (List.map Id.of_string ["Init"; "Ltac2"])) + +let std_prefix = + MPfile (DirPath.make (List.map Id.of_string ["Std"; "Ltac2"])) + +let ltac1_prefix = + MPfile (DirPath.make (List.map Id.of_string ["Ltac1"; "Ltac2"])) + +(** Generic arguments *) + +let wit_ltac2 = Genarg.make0 "ltac2:value" +let wit_ltac2_quotation = Genarg.make0 "ltac2:quotation" +let () = Geninterp.register_val0 wit_ltac2 None +let () = Geninterp.register_val0 wit_ltac2_quotation None + +let is_constructor qid = + let (_, id) = repr_qualid qid in + let id = Id.to_string id in + assert (String.length id > 0); + match id with + | "true" | "false" -> true (* built-in constructors *) + | _ -> + match id.[0] with + | 'A'..'Z' -> true + | _ -> false diff --git a/vendor/Ltac2/src/tac2env.mli b/vendor/Ltac2/src/tac2env.mli new file mode 100644 index 0000000000..c7e87c5432 --- /dev/null +++ b/vendor/Ltac2/src/tac2env.mli @@ -0,0 +1,146 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* global_data -> unit +val interp_global : ltac_constant -> global_data + +(** {5 Toplevel definition of types} *) + +val define_type : type_constant -> glb_quant_typedef -> unit +val interp_type : type_constant -> glb_quant_typedef + +(** {5 Toplevel definition of algebraic constructors} *) + +type constructor_data = { + cdata_prms : int; + (** Type parameters *) + cdata_type : type_constant; + (** Inductive definition to which the constructor pertains *) + cdata_args : int glb_typexpr list; + (** Types of the constructor arguments *) + cdata_indx : int option; + (** Index of the constructor in the ADT. Numbering is duplicated between + argumentless and argument-using constructors, e.g. in type ['a option] + [None] and [Some] have both index 0. This field is empty whenever the + constructor is a member of an open type. *) +} + +val define_constructor : ltac_constructor -> constructor_data -> unit +val interp_constructor : ltac_constructor -> constructor_data + +(** {5 Toplevel definition of projections} *) + +type projection_data = { + pdata_prms : int; + (** Type parameters *) + pdata_type : type_constant; + (** Record definition to which the projection pertains *) + pdata_ptyp : int glb_typexpr; + (** Type of the projection *) + pdata_mutb : bool; + (** Whether the field is mutable *) + pdata_indx : int; + (** Index of the projection *) +} + +val define_projection : ltac_projection -> projection_data -> unit +val interp_projection : ltac_projection -> projection_data + +(** {5 Toplevel definition of aliases} *) + +val define_alias : ltac_constant -> raw_tacexpr -> unit +val interp_alias : ltac_constant -> raw_tacexpr + +(** {5 Name management} *) + +val push_ltac : visibility -> full_path -> tacref -> unit +val locate_ltac : qualid -> tacref +val locate_extended_all_ltac : qualid -> tacref list +val shortest_qualid_of_ltac : tacref -> qualid + +val push_constructor : visibility -> full_path -> ltac_constructor -> unit +val locate_constructor : qualid -> ltac_constructor +val locate_extended_all_constructor : qualid -> ltac_constructor list +val shortest_qualid_of_constructor : ltac_constructor -> qualid + +val push_type : visibility -> full_path -> type_constant -> unit +val locate_type : qualid -> type_constant +val locate_extended_all_type : qualid -> type_constant list +val shortest_qualid_of_type : ?loc:Loc.t -> type_constant -> qualid + +val push_projection : visibility -> full_path -> ltac_projection -> unit +val locate_projection : qualid -> ltac_projection +val locate_extended_all_projection : qualid -> ltac_projection list +val shortest_qualid_of_projection : ltac_projection -> qualid + +(** {5 Toplevel definitions of ML tactics} *) + +(** This state is not part of the summary, contrarily to the ones above. It is + intended to be used from ML plugins to register ML-side functions. *) + +val define_primitive : ml_tactic_name -> closure -> unit +val interp_primitive : ml_tactic_name -> closure + +(** {5 ML primitive types} *) + +type 'a or_glb_tacexpr = +| GlbVal of 'a +| GlbTacexpr of glb_tacexpr + +type ('a, 'b, 'r) intern_fun = Genintern.glob_sign -> 'a -> 'b * 'r glb_typexpr + +type environment = { + env_ist : valexpr Id.Map.t; +} + +type ('a, 'b) ml_object = { + ml_intern : 'r. (raw_tacexpr, glb_tacexpr, 'r) intern_fun -> ('a, 'b or_glb_tacexpr, 'r) intern_fun; + ml_subst : Mod_subst.substitution -> 'b -> 'b; + ml_interp : environment -> 'b -> valexpr Proofview.tactic; + ml_print : Environ.env -> 'b -> Pp.t; +} + +val define_ml_object : ('a, 'b) Tac2dyn.Arg.tag -> ('a, 'b) ml_object -> unit +val interp_ml_object : ('a, 'b) Tac2dyn.Arg.tag -> ('a, 'b) ml_object + +(** {5 Absolute paths} *) + +val coq_prefix : ModPath.t +(** Path where primitive datatypes are defined in Ltac2 plugin. *) + +val std_prefix : ModPath.t +(** Path where Ltac-specific datatypes are defined in Ltac2 plugin. *) + +val ltac1_prefix : ModPath.t +(** Path where the Ltac1 legacy FFI is defined. *) + +(** {5 Generic arguments} *) + +val wit_ltac2 : (raw_tacexpr, glb_tacexpr, Util.Empty.t) genarg_type +val wit_ltac2_quotation : (Id.t Loc.located, Id.t, Util.Empty.t) genarg_type + +(** {5 Helper functions} *) + +val is_constructor : qualid -> bool diff --git a/vendor/Ltac2/src/tac2expr.mli b/vendor/Ltac2/src/tac2expr.mli new file mode 100644 index 0000000000..1069d0bfa3 --- /dev/null +++ b/vendor/Ltac2/src/tac2expr.mli @@ -0,0 +1,190 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* raw_tacexpr_r + +and raw_tacexpr = raw_tacexpr_r CAst.t + +and raw_taccase = raw_patexpr * raw_tacexpr + +and raw_recexpr = (ltac_projection or_relid * raw_tacexpr) list + +type case_info = type_constant or_tuple + +type 'a open_match = { + opn_match : 'a; + opn_branch : (Name.t * Name.t array * 'a) KNmap.t; + (** Invariant: should not be empty *) + opn_default : Name.t * 'a; +} + +type glb_tacexpr = +| GTacAtm of atom +| GTacVar of Id.t +| GTacRef of ltac_constant +| GTacFun of Name.t list * glb_tacexpr +| GTacApp of glb_tacexpr * glb_tacexpr list +| GTacLet of rec_flag * (Name.t * glb_tacexpr) list * glb_tacexpr +| GTacCst of case_info * int * glb_tacexpr list +| GTacCse of glb_tacexpr * case_info * glb_tacexpr array * (Name.t array * glb_tacexpr) array +| GTacPrj of type_constant * glb_tacexpr * int +| GTacSet of type_constant * glb_tacexpr * int * glb_tacexpr +| GTacOpn of ltac_constructor * glb_tacexpr list +| GTacWth of glb_tacexpr open_match +| GTacExt : (_, 'a) Tac2dyn.Arg.tag * 'a -> glb_tacexpr +| GTacPrm of ml_tactic_name * glb_tacexpr list + +(** {5 Parsing & Printing} *) + +type exp_level = +| E5 +| E4 +| E3 +| E2 +| E1 +| E0 + +type sexpr = +| SexprStr of string CAst.t +| SexprInt of int CAst.t +| SexprRec of Loc.t * Id.t option CAst.t * sexpr list + +(** {5 Toplevel statements} *) + +type strexpr = +| StrVal of mutable_flag * rec_flag * (Names.lname * raw_tacexpr) list + (** Term definition *) +| StrTyp of rec_flag * (qualid * redef_flag * raw_quant_typedef) list + (** Type definition *) +| StrPrm of Names.lident * raw_typexpr * ml_tactic_name + (** External definition *) +| StrSyn of sexpr list * int option * raw_tacexpr + (** Syntactic extensions *) +| StrMut of qualid * raw_tacexpr + (** Redefinition of mutable globals *) +| StrRun of raw_tacexpr + (** Toplevel evaluation of an expression *) + +(** {5 Dynamic semantics} *) + +(** Values are represented in a way similar to OCaml, i.e. they constrast + immediate integers (integers, constructors without arguments) and structured + blocks (tuples, arrays, constructors with arguments), as well as a few other + base cases, namely closures, strings, named constructors, and dynamic type + coming from the Coq implementation. *) + +type tag = int + +type frame = +| FrLtac of ltac_constant +| FrAnon of glb_tacexpr +| FrPrim of ml_tactic_name +| FrExtn : ('a, 'b) Tac2dyn.Arg.tag * 'b -> frame + +type backtrace = frame list diff --git a/vendor/Ltac2/src/tac2extffi.ml b/vendor/Ltac2/src/tac2extffi.ml new file mode 100644 index 0000000000..315c970f9e --- /dev/null +++ b/vendor/Ltac2/src/tac2extffi.ml @@ -0,0 +1,40 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* assert false) f + +(** More ML representations *) + +let to_qhyp v = match Value.to_block v with +| (0, [| i |]) -> AnonHyp (Value.to_int i) +| (1, [| id |]) -> NamedHyp (Value.to_ident id) +| _ -> assert false + +let qhyp = make_to_repr to_qhyp + +let to_bindings = function +| ValInt 0 -> NoBindings +| ValBlk (0, [| vl |]) -> + ImplicitBindings (Value.to_list Value.to_constr vl) +| ValBlk (1, [| vl |]) -> + ExplicitBindings ((Value.to_list (fun p -> to_pair to_qhyp Value.to_constr p) vl)) +| _ -> assert false + +let bindings = make_to_repr to_bindings + +let to_constr_with_bindings v = match Value.to_tuple v with +| [| c; bnd |] -> (Value.to_constr c, to_bindings bnd) +| _ -> assert false + +let constr_with_bindings = make_to_repr to_constr_with_bindings diff --git a/vendor/Ltac2/src/tac2extffi.mli b/vendor/Ltac2/src/tac2extffi.mli new file mode 100644 index 0000000000..f5251c3d0d --- /dev/null +++ b/vendor/Ltac2/src/tac2extffi.mli @@ -0,0 +1,16 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* 'a Proofview.tactic) arity0 +| AddAty : ('a, 'b) arity0 -> ('a, 'a -> 'b) arity0 + +type valexpr = +| ValInt of int + (** Immediate integers *) +| ValBlk of tag * valexpr array + (** Structured blocks *) +| ValStr of Bytes.t + (** Strings *) +| ValCls of closure + (** Closures *) +| ValOpn of KerName.t * valexpr array + (** Open constructors *) +| ValExt : 'a Tac2dyn.Val.tag * 'a -> valexpr + (** Arbitrary data *) + +and closure = MLTactic : (valexpr, 'v) arity0 * 'v -> closure + +let arity_one = OneAty +let arity_suc a = AddAty a + +type 'a arity = (valexpr, 'a) arity0 + +let mk_closure arity f = MLTactic (arity, f) + +module Valexpr = +struct + +type t = valexpr + +let is_int = function +| ValInt _ -> true +| ValBlk _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ -> false + +let tag v = match v with +| ValBlk (n, _) -> n +| ValInt _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ -> + CErrors.anomaly (Pp.str "Unexpected value shape") + +let field v n = match v with +| ValBlk (_, v) -> v.(n) +| ValInt _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ -> + CErrors.anomaly (Pp.str "Unexpected value shape") + +let set_field v n w = match v with +| ValBlk (_, v) -> v.(n) <- w +| ValInt _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ -> + CErrors.anomaly (Pp.str "Unexpected value shape") + +let make_block tag v = ValBlk (tag, v) +let make_int n = ValInt n + +end + +type 'a repr = { + r_of : 'a -> valexpr; + r_to : valexpr -> 'a; + r_id : bool; +} + +let repr_of r x = r.r_of x +let repr_to r x = r.r_to x + +let make_repr r_of r_to = { r_of; r_to; r_id = false; } + +(** Dynamic tags *) + +let val_exn = Val.create "exn" +let val_constr = Val.create "constr" +let val_ident = Val.create "ident" +let val_pattern = Val.create "pattern" +let val_pp = Val.create "pp" +let val_sort = Val.create "sort" +let val_cast = Val.create "cast" +let val_inductive = Val.create "inductive" +let val_constant = Val.create "constant" +let val_constructor = Val.create "constructor" +let val_projection = Val.create "projection" +let val_case = Val.create "case" +let val_univ = Val.create "universe" +let val_free : Names.Id.Set.t Val.tag = Val.create "free" +let val_ltac1 : Geninterp.Val.t Val.tag = Val.create "ltac1" + +let extract_val (type a) (type b) (tag : a Val.tag) (tag' : b Val.tag) (v : b) : a = +match Val.eq tag tag' with +| None -> assert false +| Some Refl -> v + +(** Exception *) + +exception LtacError of KerName.t * valexpr array + +(** Conversion functions *) + +let valexpr = { + r_of = (fun obj -> obj); + r_to = (fun obj -> obj); + r_id = true; +} + +let of_unit () = ValInt 0 + +let to_unit = function +| ValInt 0 -> () +| _ -> assert false + +let unit = { + r_of = of_unit; + r_to = to_unit; + r_id = false; +} + +let of_int n = ValInt n +let to_int = function +| ValInt n -> n +| _ -> assert false + +let int = { + r_of = of_int; + r_to = to_int; + r_id = false; +} + +let of_bool b = if b then ValInt 0 else ValInt 1 + +let to_bool = function +| ValInt 0 -> true +| ValInt 1 -> false +| _ -> assert false + +let bool = { + r_of = of_bool; + r_to = to_bool; + r_id = false; +} + +let of_char n = ValInt (Char.code n) +let to_char = function +| ValInt n -> Char.chr n +| _ -> assert false + +let char = { + r_of = of_char; + r_to = to_char; + r_id = false; +} + +let of_string s = ValStr s +let to_string = function +| ValStr s -> s +| _ -> assert false + +let string = { + r_of = of_string; + r_to = to_string; + r_id = false; +} + +let rec of_list f = function +| [] -> ValInt 0 +| x :: l -> ValBlk (0, [| f x; of_list f l |]) + +let rec to_list f = function +| ValInt 0 -> [] +| ValBlk (0, [|v; vl|]) -> f v :: to_list f vl +| _ -> assert false + +let list r = { + r_of = (fun l -> of_list r.r_of l); + r_to = (fun l -> to_list r.r_to l); + r_id = false; +} + +let of_closure cls = ValCls cls + +let to_closure = function +| ValCls cls -> cls +| ValExt _ | ValInt _ | ValBlk _ | ValStr _ | ValOpn _ -> assert false + +let closure = { + r_of = of_closure; + r_to = to_closure; + r_id = false; +} + +let of_ext tag c = + ValExt (tag, c) + +let to_ext tag = function +| ValExt (tag', e) -> extract_val tag tag' e +| _ -> assert false + +let repr_ext tag = { + r_of = (fun e -> of_ext tag e); + r_to = (fun e -> to_ext tag e); + r_id = false; +} + +let of_constr c = of_ext val_constr c +let to_constr c = to_ext val_constr c +let constr = repr_ext val_constr + +let of_ident c = of_ext val_ident c +let to_ident c = to_ext val_ident c +let ident = repr_ext val_ident + +let of_pattern c = of_ext val_pattern c +let to_pattern c = to_ext val_pattern c +let pattern = repr_ext val_pattern + +let internal_err = + let open Names in + let coq_prefix = + MPfile (DirPath.make (List.map Id.of_string ["Init"; "Ltac2"])) + in + KerName.make coq_prefix (Label.of_id (Id.of_string "Internal")) + +(** FIXME: handle backtrace in Ltac2 exceptions *) +let of_exn c = match fst c with +| LtacError (kn, c) -> ValOpn (kn, c) +| _ -> ValOpn (internal_err, [|of_ext val_exn c|]) + +let to_exn c = match c with +| ValOpn (kn, c) -> + if Names.KerName.equal kn internal_err then + to_ext val_exn c.(0) + else + (LtacError (kn, c), Exninfo.null) +| _ -> assert false + +let exn = { + r_of = of_exn; + r_to = to_exn; + r_id = false; +} + +let of_option f = function +| None -> ValInt 0 +| Some c -> ValBlk (0, [|f c|]) + +let to_option f = function +| ValInt 0 -> None +| ValBlk (0, [|c|]) -> Some (f c) +| _ -> assert false + +let option r = { + r_of = (fun l -> of_option r.r_of l); + r_to = (fun l -> to_option r.r_to l); + r_id = false; +} + +let of_pp c = of_ext val_pp c +let to_pp c = to_ext val_pp c +let pp = repr_ext val_pp + +let of_tuple cl = ValBlk (0, cl) +let to_tuple = function +| ValBlk (0, cl) -> cl +| _ -> assert false + +let of_pair f g (x, y) = ValBlk (0, [|f x; g y|]) +let to_pair f g = function +| ValBlk (0, [|x; y|]) -> (f x, g y) +| _ -> assert false +let pair r0 r1 = { + r_of = (fun p -> of_pair r0.r_of r1.r_of p); + r_to = (fun p -> to_pair r0.r_to r1.r_to p); + r_id = false; +} + +let of_array f vl = ValBlk (0, Array.map f vl) +let to_array f = function +| ValBlk (0, vl) -> Array.map f vl +| _ -> assert false +let array r = { + r_of = (fun l -> of_array r.r_of l); + r_to = (fun l -> to_array r.r_to l); + r_id = false; +} + +let of_block (n, args) = ValBlk (n, args) +let to_block = function +| ValBlk (n, args) -> (n, args) +| _ -> assert false + +let block = { + r_of = of_block; + r_to = to_block; + r_id = false; +} + +let of_open (kn, args) = ValOpn (kn, args) + +let to_open = function +| ValOpn (kn, args) -> (kn, args) +| _ -> assert false + +let open_ = { + r_of = of_open; + r_to = to_open; + r_id = false; +} + +let of_constant c = of_ext val_constant c +let to_constant c = to_ext val_constant c +let constant = repr_ext val_constant + +let of_reference = function +| VarRef id -> ValBlk (0, [| of_ident id |]) +| ConstRef cst -> ValBlk (1, [| of_constant cst |]) +| IndRef ind -> ValBlk (2, [| of_ext val_inductive ind |]) +| ConstructRef cstr -> ValBlk (3, [| of_ext val_constructor cstr |]) + +let to_reference = function +| ValBlk (0, [| id |]) -> VarRef (to_ident id) +| ValBlk (1, [| cst |]) -> ConstRef (to_constant cst) +| ValBlk (2, [| ind |]) -> IndRef (to_ext val_inductive ind) +| ValBlk (3, [| cstr |]) -> ConstructRef (to_ext val_constructor cstr) +| _ -> assert false + +let reference = { + r_of = of_reference; + r_to = to_reference; + r_id = false; +} + +type ('a, 'b) fun1 = closure + +let fun1 (r0 : 'a repr) (r1 : 'b repr) : ('a, 'b) fun1 repr = closure +let to_fun1 r0 r1 f = to_closure f + +let rec apply : type a. a arity -> a -> valexpr list -> valexpr Proofview.tactic = + fun arity f args -> match args, arity with + | [], arity -> Proofview.tclUNIT (ValCls (MLTactic (arity, f))) + (* A few hardcoded cases for efficiency *) + | [a0], OneAty -> f a0 + | [a0; a1], AddAty OneAty -> f a0 a1 + | [a0; a1; a2], AddAty (AddAty OneAty) -> f a0 a1 a2 + | [a0; a1; a2; a3], AddAty (AddAty (AddAty OneAty)) -> f a0 a1 a2 a3 + (* Generic cases *) + | a :: args, OneAty -> + f a >>= fun f -> + let MLTactic (arity, f) = to_closure f in + apply arity f args + | a :: args, AddAty arity -> + apply arity (f a) args + +let apply (MLTactic (arity, f)) args = apply arity f args + +type n_closure = +| NClosure : 'a arity * (valexpr list -> 'a) -> n_closure + +let rec abstract n f = + if Int.equal n 1 then NClosure (OneAty, fun accu v -> f (List.rev (v :: accu))) + else + let NClosure (arity, fe) = abstract (n - 1) f in + NClosure (AddAty arity, fun accu v -> fe (v :: accu)) + +let abstract n f = + let () = assert (n > 0) in + let NClosure (arity, f) = abstract n f in + MLTactic (arity, f []) + +let app_fun1 cls r0 r1 x = + apply cls [r0.r_of x] >>= fun v -> Proofview.tclUNIT (r1.r_to v) diff --git a/vendor/Ltac2/src/tac2ffi.mli b/vendor/Ltac2/src/tac2ffi.mli new file mode 100644 index 0000000000..bfc93d99e6 --- /dev/null +++ b/vendor/Ltac2/src/tac2ffi.mli @@ -0,0 +1,189 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* valexpr + (** Arbitrary data *) + +type 'a arity + +val arity_one : (valexpr -> valexpr Proofview.tactic) arity +val arity_suc : 'a arity -> (valexpr -> 'a) arity + +val mk_closure : 'v arity -> 'v -> closure + +module Valexpr : +sig + type t = valexpr + val is_int : t -> bool + val tag : t -> int + val field : t -> int -> t + val set_field : t -> int -> t -> unit + val make_block : int -> t array -> t + val make_int : int -> t +end + +(** {5 Ltac2 FFI} *) + +type 'a repr + +val repr_of : 'a repr -> 'a -> valexpr +val repr_to : 'a repr -> valexpr -> 'a + +val make_repr : ('a -> valexpr) -> (valexpr -> 'a) -> 'a repr + +(** These functions allow to convert back and forth between OCaml and Ltac2 + data representation. The [to_*] functions raise an anomaly whenever the data + has not expected shape. *) + +val of_unit : unit -> valexpr +val to_unit : valexpr -> unit +val unit : unit repr + +val of_int : int -> valexpr +val to_int : valexpr -> int +val int : int repr + +val of_bool : bool -> valexpr +val to_bool : valexpr -> bool +val bool : bool repr + +val of_char : char -> valexpr +val to_char : valexpr -> char +val char : char repr + +val of_string : Bytes.t -> valexpr +val to_string : valexpr -> Bytes.t +val string : Bytes.t repr + +val of_list : ('a -> valexpr) -> 'a list -> valexpr +val to_list : (valexpr -> 'a) -> valexpr -> 'a list +val list : 'a repr -> 'a list repr + +val of_constr : EConstr.t -> valexpr +val to_constr : valexpr -> EConstr.t +val constr : EConstr.t repr + +val of_exn : Exninfo.iexn -> valexpr +val to_exn : valexpr -> Exninfo.iexn +val exn : Exninfo.iexn repr + +val of_ident : Id.t -> valexpr +val to_ident : valexpr -> Id.t +val ident : Id.t repr + +val of_closure : closure -> valexpr +val to_closure : valexpr -> closure +val closure : closure repr + +val of_block : (int * valexpr array) -> valexpr +val to_block : valexpr -> (int * valexpr array) +val block : (int * valexpr array) repr + +val of_array : ('a -> valexpr) -> 'a array -> valexpr +val to_array : (valexpr -> 'a) -> valexpr -> 'a array +val array : 'a repr -> 'a array repr + +val of_tuple : valexpr array -> valexpr +val to_tuple : valexpr -> valexpr array + +val of_pair : ('a -> valexpr) -> ('b -> valexpr) -> 'a * 'b -> valexpr +val to_pair : (valexpr -> 'a) -> (valexpr -> 'b) -> valexpr -> 'a * 'b +val pair : 'a repr -> 'b repr -> ('a * 'b) repr + +val of_option : ('a -> valexpr) -> 'a option -> valexpr +val to_option : (valexpr -> 'a) -> valexpr -> 'a option +val option : 'a repr -> 'a option repr + +val of_pattern : Pattern.constr_pattern -> valexpr +val to_pattern : valexpr -> Pattern.constr_pattern +val pattern : Pattern.constr_pattern repr + +val of_pp : Pp.t -> valexpr +val to_pp : valexpr -> Pp.t +val pp : Pp.t repr + +val of_constant : Constant.t -> valexpr +val to_constant : valexpr -> Constant.t +val constant : Constant.t repr + +val of_reference : GlobRef.t -> valexpr +val to_reference : valexpr -> GlobRef.t +val reference : GlobRef.t repr + +val of_ext : 'a Val.tag -> 'a -> valexpr +val to_ext : 'a Val.tag -> valexpr -> 'a +val repr_ext : 'a Val.tag -> 'a repr + +val of_open : KerName.t * valexpr array -> valexpr +val to_open : valexpr -> KerName.t * valexpr array +val open_ : (KerName.t * valexpr array) repr + +type ('a, 'b) fun1 + +val app_fun1 : ('a, 'b) fun1 -> 'a repr -> 'b repr -> 'a -> 'b Proofview.tactic + +val to_fun1 : 'a repr -> 'b repr -> valexpr -> ('a, 'b) fun1 +val fun1 : 'a repr -> 'b repr -> ('a, 'b) fun1 repr + +val valexpr : valexpr repr + +(** {5 Dynamic tags} *) + +val val_constr : EConstr.t Val.tag +val val_ident : Id.t Val.tag +val val_pattern : Pattern.constr_pattern Val.tag +val val_pp : Pp.t Val.tag +val val_sort : ESorts.t Val.tag +val val_cast : Constr.cast_kind Val.tag +val val_inductive : inductive Val.tag +val val_constant : Constant.t Val.tag +val val_constructor : constructor Val.tag +val val_projection : Projection.t Val.tag +val val_case : Constr.case_info Val.tag +val val_univ : Univ.Level.t Val.tag +val val_free : Id.Set.t Val.tag +val val_ltac1 : Geninterp.Val.t Val.tag + +val val_exn : Exninfo.iexn Tac2dyn.Val.tag +(** Toplevel representation of OCaml exceptions. Invariant: no [LtacError] + should be put into a value with tag [val_exn]. *) + +(** Closures *) + +val apply : closure -> valexpr list -> valexpr Proofview.tactic +(** Given a closure, apply it to some arguments. Handling of argument mismatches + is done automatically, i.e. in case of over or under-application. *) + +val abstract : int -> (valexpr list -> valexpr Proofview.tactic) -> closure +(** Turn a fixed-arity function into a closure. The inner function is guaranteed + to be applied to a list whose size is the integer argument. *) + +(** Exception *) + +exception LtacError of KerName.t * valexpr array +(** Ltac2-defined exceptions seen from OCaml side *) diff --git a/vendor/Ltac2/src/tac2intern.ml b/vendor/Ltac2/src/tac2intern.ml new file mode 100644 index 0000000000..de99fb167f --- /dev/null +++ b/vendor/Ltac2/src/tac2intern.ml @@ -0,0 +1,1545 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* elt -> bool +val create : unit -> 'a t +val fresh : 'a t -> elt +val find : elt -> 'a t -> (elt * 'a option) +val union : elt -> elt -> 'a t -> unit +val set : elt -> 'a -> 'a t -> unit +module Map : +sig + type key = elt + type +'a t + val empty : 'a t + val add : key -> 'a -> 'a t -> 'a t + val mem : key -> 'a t -> bool + val find : key -> 'a t -> 'a + val exists : (key -> 'a -> bool) -> 'a t -> bool +end +end += +struct +type elt = int +let equal = Int.equal +module Map = Int.Map + +type 'a node = +| Canon of int * 'a option +| Equiv of elt + +type 'a t = { + mutable uf_data : 'a node array; + mutable uf_size : int; +} + +let resize p = + if Int.equal (Array.length p.uf_data) p.uf_size then begin + let nsize = 2 * p.uf_size + 1 in + let v = Array.make nsize (Equiv 0) in + Array.blit p.uf_data 0 v 0 (Array.length p.uf_data); + p.uf_data <- v; + end + +let create () = { uf_data = [||]; uf_size = 0 } + +let fresh p = + resize p; + let n = p.uf_size in + p.uf_data.(n) <- (Canon (1, None)); + p.uf_size <- n + 1; + n + +let rec lookup n p = + let node = Array.get p.uf_data n in + match node with + | Canon (size, v) -> n, size, v + | Equiv y -> + let ((z, _, _) as res) = lookup y p in + if not (Int.equal z y) then Array.set p.uf_data n (Equiv z); + res + +let find n p = + let (x, _, v) = lookup n p in (x, v) + +let union x y p = + let ((x, size1, _) as xcan) = lookup x p in + let ((y, size2, _) as ycan) = lookup y p in + let xcan, ycan = if size1 < size2 then xcan, ycan else ycan, xcan in + let x, _, xnode = xcan in + let y, _, ynode = ycan in + assert (Option.is_empty xnode); + assert (Option.is_empty ynode); + p.uf_data.(x) <- Equiv y; + p.uf_data.(y) <- Canon (size1 + size2, None) + +let set x v p = + let (x, s, v') = lookup x p in + assert (Option.is_empty v'); + p.uf_data.(x) <- Canon (s, Some v) + +end + +type mix_var = +| GVar of UF.elt +| LVar of int + +type mix_type_scheme = int * mix_var glb_typexpr + +type environment = { + env_var : mix_type_scheme Id.Map.t; + (** Type schemes of bound variables *) + env_cst : UF.elt glb_typexpr UF.t; + (** Unification state *) + env_als : UF.elt Id.Map.t ref; + (** Map user-facing type variables to unification variables *) + env_opn : bool; + (** Accept unbound type variables *) + env_rec : (KerName.t * int) Id.Map.t; + (** Recursive type definitions *) + env_str : bool; + (** True iff in strict mode *) +} + +let empty_env () = { + env_var = Id.Map.empty; + env_cst = UF.create (); + env_als = ref Id.Map.empty; + env_opn = true; + env_rec = Id.Map.empty; + env_str = true; +} + +let env_name env = + (* Generate names according to a provided environment *) + let mk num = + let base = num mod 26 in + let rem = num / 26 in + let name = String.make 1 (Char.chr (97 + base)) in + let suff = if Int.equal rem 0 then "" else string_of_int rem in + let name = name ^ suff in + name + in + let fold id elt acc = UF.Map.add elt (Id.to_string id) acc in + let vars = Id.Map.fold fold env.env_als.contents UF.Map.empty in + let vars = ref vars in + let rec fresh n = + let name = mk n in + if UF.Map.exists (fun _ name' -> String.equal name name') !vars then fresh (succ n) + else name + in + fun n -> + if UF.Map.mem n !vars then UF.Map.find n !vars + else + let ans = fresh 0 in + let () = vars := UF.Map.add n ans !vars in + ans + +let ltac2_env : environment Genintern.Store.field = + Genintern.Store.field () + +let drop_ltac2_env store = + Genintern.Store.remove store ltac2_env + +let fresh_id env = UF.fresh env.env_cst + +let get_alias {loc;v=id} env = + try Id.Map.find id env.env_als.contents + with Not_found -> + if env.env_opn then + let n = fresh_id env in + let () = env.env_als := Id.Map.add id n env.env_als.contents in + n + else user_err ?loc (str "Unbound type parameter " ++ Id.print id) + +let push_name id t env = match id with +| Anonymous -> env +| Name id -> { env with env_var = Id.Map.add id t env.env_var } + +let error_nargs_mismatch ?loc kn nargs nfound = + let cstr = Tac2env.shortest_qualid_of_constructor kn in + user_err ?loc (str "Constructor " ++ pr_qualid cstr ++ str " expects " ++ + int nargs ++ str " arguments, but is applied to " ++ int nfound ++ + str " arguments") + +let error_nparams_mismatch ?loc nargs nfound = + user_err ?loc (str "Type expects " ++ int nargs ++ + str " arguments, but is applied to " ++ int nfound ++ + str " arguments") + +let rec subst_type subst (t : 'a glb_typexpr) = match t with +| GTypVar id -> subst id +| GTypArrow (t1, t2) -> GTypArrow (subst_type subst t1, subst_type subst t2) +| GTypRef (qid, args) -> + GTypRef (qid, List.map (fun t -> subst_type subst t) args) + +let rec intern_type env ({loc;v=t} : raw_typexpr) : UF.elt glb_typexpr = match t with +| CTypVar (Name id) -> GTypVar (get_alias (CAst.make ?loc id) env) +| CTypVar Anonymous -> GTypVar (fresh_id env) +| CTypRef (rel, args) -> + let (kn, nparams) = match rel with + | RelId qid -> + let id = qualid_basename qid in + if qualid_is_ident qid && Id.Map.mem id env.env_rec then + let (kn, n) = Id.Map.find id env.env_rec in + (Other kn, n) + else + let kn = + try Tac2env.locate_type qid + with Not_found -> + user_err ?loc (str "Unbound type constructor " ++ pr_qualid qid) + in + let (nparams, _) = Tac2env.interp_type kn in + (Other kn, nparams) + | AbsKn (Other kn) -> + let (nparams, _) = Tac2env.interp_type kn in + (Other kn, nparams) + | AbsKn (Tuple n) -> + (Tuple n, n) + in + let nargs = List.length args in + let () = + if not (Int.equal nparams nargs) then + let qid = match rel with + | RelId lid -> lid + | AbsKn (Other kn) -> shortest_qualid_of_type ?loc kn + | AbsKn (Tuple _) -> assert false + in + user_err ?loc (strbrk "The type constructor " ++ pr_qualid qid ++ + strbrk " expects " ++ int nparams ++ strbrk " argument(s), but is here \ + applied to " ++ int nargs ++ strbrk "argument(s)") + in + GTypRef (kn, List.map (fun t -> intern_type env t) args) +| CTypArrow (t1, t2) -> GTypArrow (intern_type env t1, intern_type env t2) + +let fresh_type_scheme env (t : type_scheme) : UF.elt glb_typexpr = + let (n, t) = t in + let subst = Array.init n (fun _ -> fresh_id env) in + let substf i = GTypVar subst.(i) in + subst_type substf t + +let fresh_mix_type_scheme env (t : mix_type_scheme) : UF.elt glb_typexpr = + let (n, t) = t in + let subst = Array.init n (fun _ -> fresh_id env) in + let substf = function + | LVar i -> GTypVar subst.(i) + | GVar n -> GTypVar n + in + subst_type substf t + +let fresh_reftype env (kn : KerName.t or_tuple) = + let n = match kn with + | Other kn -> fst (Tac2env.interp_type kn) + | Tuple n -> n + in + let subst = Array.init n (fun _ -> fresh_id env) in + let t = GTypRef (kn, Array.map_to_list (fun i -> GTypVar i) subst) in + (subst, t) + +(** First-order unification algorithm *) +let is_unfoldable kn = match snd (Tac2env.interp_type kn) with +| GTydDef (Some _) -> true +| GTydDef None | GTydAlg _ | GTydRec _ | GTydOpn -> false + +let unfold env kn args = + let (nparams, def) = Tac2env.interp_type kn in + let def = match def with + | GTydDef (Some t) -> t + | _ -> assert false + in + let args = Array.of_list args in + let subst n = args.(n) in + subst_type subst def + +(** View function, allows to ensure head normal forms *) +let rec kind env t = match t with +| GTypVar id -> + let (id, v) = UF.find id env.env_cst in + begin match v with + | None -> GTypVar id + | Some t -> kind env t + end +| GTypRef (Other kn, tl) -> + if is_unfoldable kn then kind env (unfold env kn tl) else t +| GTypArrow _ | GTypRef (Tuple _, _) -> t + +(** Normalize unification variables without unfolding type aliases *) +let rec nf env t = match t with +| GTypVar id -> + let (id, v) = UF.find id env.env_cst in + begin match v with + | None -> GTypVar id + | Some t -> nf env t + end +| GTypRef (kn, tl) -> + let tl = List.map (fun t -> nf env t) tl in + GTypRef (kn, tl) +| GTypArrow (t, u) -> + let t = nf env t in + let u = nf env u in + GTypArrow (t, u) + +let pr_glbtype env t = + let t = nf env t in + let name = env_name env in + pr_glbtype name t + +exception Occur + +let rec occur_check env id t = match kind env t with +| GTypVar id' -> if UF.equal id id' then raise Occur +| GTypArrow (t1, t2) -> + let () = occur_check env id t1 in + occur_check env id t2 +| GTypRef (kn, tl) -> + List.iter (fun t -> occur_check env id t) tl + +exception CannotUnify of UF.elt glb_typexpr * UF.elt glb_typexpr + +let unify_var env id t = match kind env t with +| GTypVar id' -> + if not (UF.equal id id') then UF.union id id' env.env_cst +| GTypArrow _ | GTypRef _ -> + try + let () = occur_check env id t in + UF.set id t env.env_cst + with Occur -> raise (CannotUnify (GTypVar id, t)) + +let eq_or_tuple eq t1 t2 = match t1, t2 with +| Tuple n1, Tuple n2 -> Int.equal n1 n2 +| Other o1, Other o2 -> eq o1 o2 +| _ -> false + +let rec unify0 env t1 t2 = match kind env t1, kind env t2 with +| GTypVar id, t | t, GTypVar id -> + unify_var env id t +| GTypArrow (t1, u1), GTypArrow (t2, u2) -> + let () = unify0 env t1 t2 in + unify0 env u1 u2 +| GTypRef (kn1, tl1), GTypRef (kn2, tl2) -> + if eq_or_tuple KerName.equal kn1 kn2 then + List.iter2 (fun t1 t2 -> unify0 env t1 t2) tl1 tl2 + else raise (CannotUnify (t1, t2)) +| _ -> raise (CannotUnify (t1, t2)) + +let unify ?loc env t1 t2 = + try unify0 env t1 t2 + with CannotUnify (u1, u2) -> + user_err ?loc (str "This expression has type" ++ spc () ++ pr_glbtype env t1 ++ + spc () ++ str "but an expression was expected of type" ++ spc () ++ pr_glbtype env t2) + +let unify_arrow ?loc env ft args = + let ft0 = ft in + let rec iter ft args is_fun = match kind env ft, args with + | t, [] -> t + | GTypArrow (t1, ft), (loc, t2) :: args -> + let () = unify ?loc env t2 t1 in + iter ft args true + | GTypVar id, (_, t) :: args -> + let ft = GTypVar (fresh_id env) in + let () = unify_var env id (GTypArrow (t, ft)) in + iter ft args true + | GTypRef _, _ :: _ -> + if is_fun then + user_err ?loc (str "This function has type" ++ spc () ++ pr_glbtype env ft0 ++ + spc () ++ str "and is applied to too many arguments") + else + user_err ?loc (str "This expression has type" ++ spc () ++ pr_glbtype env ft0 ++ + spc () ++ str "and is not a function") + in + iter ft args false + +(** Term typing *) + +let is_pure_constructor kn = + match snd (Tac2env.interp_type kn) with + | GTydAlg _ | GTydOpn -> true + | GTydRec fields -> + let is_pure (_, mut, _) = not mut in + List.for_all is_pure fields + | GTydDef _ -> assert false (** Type definitions have no constructors *) + +let rec is_value = function +| GTacAtm (AtmInt _) | GTacVar _ | GTacRef _ | GTacFun _ -> true +| GTacAtm (AtmStr _) | GTacApp _ | GTacLet _ -> false +| GTacCst (Tuple _, _, el) -> List.for_all is_value el +| GTacCst (_, _, []) -> true +| GTacOpn (_, el) -> List.for_all is_value el +| GTacCst (Other kn, _, el) -> is_pure_constructor kn && List.for_all is_value el +| GTacCse _ | GTacPrj _ | GTacSet _ | GTacExt _ | GTacPrm _ +| GTacWth _ -> false + +let is_rec_rhs = function +| GTacFun _ -> true +| GTacAtm _ | GTacVar _ | GTacRef _ | GTacApp _ | GTacLet _ | GTacPrj _ +| GTacSet _ | GTacExt _ | GTacPrm _ | GTacCst _ +| GTacCse _ | GTacOpn _ | GTacWth _ -> false + +let rec fv_type f t accu = match t with +| GTypVar id -> f id accu +| GTypArrow (t1, t2) -> fv_type f t1 (fv_type f t2 accu) +| GTypRef (kn, tl) -> List.fold_left (fun accu t -> fv_type f t accu) accu tl + +let fv_env env = + let rec f id accu = match UF.find id env.env_cst with + | id, None -> UF.Map.add id () accu + | _, Some t -> fv_type f t accu + in + let fold_var id (_, t) accu = + let fmix id accu = match id with + | LVar _ -> accu + | GVar id -> f id accu + in + fv_type fmix t accu + in + let fv_var = Id.Map.fold fold_var env.env_var UF.Map.empty in + let fold_als _ id accu = f id accu in + Id.Map.fold fold_als !(env.env_als) fv_var + +let abstract_var env (t : UF.elt glb_typexpr) : mix_type_scheme = + let fv = fv_env env in + let count = ref 0 in + let vars = ref UF.Map.empty in + let rec subst id = + let (id, t) = UF.find id env.env_cst in + match t with + | None -> + if UF.Map.mem id fv then GTypVar (GVar id) + else + begin try UF.Map.find id !vars + with Not_found -> + let n = !count in + let var = GTypVar (LVar n) in + let () = incr count in + let () = vars := UF.Map.add id var !vars in + var + end + | Some t -> subst_type subst t + in + let t = subst_type subst t in + (!count, t) + +let monomorphic (t : UF.elt glb_typexpr) : mix_type_scheme = + let subst id = GTypVar (GVar id) in + (0, subst_type subst t) + +let warn_not_unit = + CWarnings.create ~name:"not-unit" ~category:"ltac" + (fun () -> strbrk "The following expression should have type unit.") + +let warn_redundant_clause = + CWarnings.create ~name:"redundant-clause" ~category:"ltac" + (fun () -> strbrk "The following clause is redundant.") + +let check_elt_unit loc env t = + let maybe_unit = match kind env t with + | GTypVar _ -> true + | GTypArrow _ -> false + | GTypRef (Tuple 0, []) -> true + | GTypRef _ -> false + in + if not maybe_unit then warn_not_unit ?loc () + +let check_elt_empty loc env t = match kind env t with +| GTypVar _ -> + user_err ?loc (str "Cannot infer an empty type for this expression") +| GTypArrow _ | GTypRef (Tuple _, _) -> + user_err ?loc (str "Type" ++ spc () ++ pr_glbtype env t ++ spc () ++ str "is not an empty type") +| GTypRef (Other kn, _) -> + let def = Tac2env.interp_type kn in + match def with + | _, GTydAlg { galg_constructors = [] } -> kn + | _ -> + user_err ?loc (str "Type" ++ spc () ++ pr_glbtype env t ++ spc () ++ str "is not an empty type") + +let check_unit ?loc t = + let env = empty_env () in + (* Should not matter, t should be closed. *) + let t = fresh_type_scheme env t in + let maybe_unit = match kind env t with + | GTypVar _ -> true + | GTypArrow _ -> false + | GTypRef (Tuple 0, []) -> true + | GTypRef _ -> false + in + if not maybe_unit then warn_not_unit ?loc () + +let check_redundant_clause = function +| [] -> () +| (p, _) :: _ -> warn_redundant_clause ?loc:p.loc () + +let get_variable0 mem var = match var with +| RelId qid -> + let id = qualid_basename qid in + if qualid_is_ident qid && mem id then ArgVar CAst.(make ?loc:qid.CAst.loc id) + else + let kn = + try Tac2env.locate_ltac qid + with Not_found -> + CErrors.user_err ?loc:qid.CAst.loc (str "Unbound value " ++ pr_qualid qid) + in + ArgArg kn +| AbsKn kn -> ArgArg kn + +let get_variable env var = + let mem id = Id.Map.mem id env.env_var in + get_variable0 mem var + +let get_constructor env var = match var with +| RelId qid -> + let c = try Some (Tac2env.locate_constructor qid) with Not_found -> None in + begin match c with + | Some knc -> Other knc + | None -> + CErrors.user_err ?loc:qid.CAst.loc (str "Unbound constructor " ++ pr_qualid qid) + end +| AbsKn knc -> knc + +let get_projection var = match var with +| RelId qid -> + let kn = try Tac2env.locate_projection qid with Not_found -> + user_err ?loc:qid.CAst.loc (pr_qualid qid ++ str " is not a projection") + in + Tac2env.interp_projection kn +| AbsKn kn -> + Tac2env.interp_projection kn + +let intern_atm env = function +| AtmInt n -> (GTacAtm (AtmInt n), GTypRef (Other t_int, [])) +| AtmStr s -> (GTacAtm (AtmStr s), GTypRef (Other t_string, [])) + +let invalid_pattern ?loc kn kn' = + let pr t = match t with + | Other kn' -> str "type " ++ pr_typref kn' + | Tuple n -> str "tuple of size " ++ int n + in + user_err ?loc (str "Invalid pattern, expected a pattern for " ++ + pr kn ++ str ", found a pattern for " ++ pr kn') (** FIXME *) + +(** Pattern view *) + +type glb_patexpr = +| GPatVar of Name.t +| GPatRef of ltac_constructor or_tuple * glb_patexpr list + +let rec intern_patexpr env {loc;v=pat} = match pat with +| CPatVar na -> GPatVar na +| CPatRef (qid, pl) -> + let kn = get_constructor env qid in + GPatRef (kn, List.map (fun p -> intern_patexpr env p) pl) +| CPatCnv (pat, ty) -> + user_err ?loc (str "Pattern not handled yet") + +type pattern_kind = +| PKind_empty +| PKind_variant of type_constant or_tuple +| PKind_open of type_constant +| PKind_any + +let get_pattern_kind env pl = match pl with +| [] -> PKind_empty +| p :: pl -> + let rec get_kind (p, _) pl = match intern_patexpr env p with + | GPatVar _ -> + begin match pl with + | [] -> PKind_any + | p :: pl -> get_kind p pl + end + | GPatRef (Other kn, pl) -> + let data = Tac2env.interp_constructor kn in + if Option.is_empty data.cdata_indx then PKind_open data.cdata_type + else PKind_variant (Other data.cdata_type) + | GPatRef (Tuple _, tp) -> PKind_variant (Tuple (List.length tp)) + in + get_kind p pl + +(** Internalization *) + +(** Used to generate a fresh tactic variable for pattern-expansion *) +let fresh_var avoid = + let bad id = + Id.Set.mem id avoid || + (try ignore (locate_ltac (qualid_of_ident id)); true with Not_found -> false) + in + Namegen.next_ident_away_from (Id.of_string "p") bad + +let add_name accu = function +| Name id -> Id.Set.add id accu +| Anonymous -> accu + +let rec ids_of_pattern accu {v=pat} = match pat with +| CPatVar Anonymous -> accu +| CPatVar (Name id) -> Id.Set.add id accu +| CPatRef (_, pl) -> + List.fold_left ids_of_pattern accu pl +| CPatCnv (pat, _) -> ids_of_pattern accu pat + +let loc_of_relid = function +| RelId {loc} -> loc +| AbsKn _ -> None + +let extract_pattern_type ({loc;v=p} as pat) = match p with +| CPatCnv (pat, ty) -> pat, Some ty +| CPatVar _ | CPatRef _ -> pat, None + +(** Expand pattern: [p => t] becomes [x => match x with p => t end] *) +let expand_pattern avoid bnd = + let fold (avoid, bnd) (pat, t) = + let na, expand = match pat.v with + | CPatVar na -> + (* Don't expand variable patterns *) + na, None + | _ -> + let id = fresh_var avoid in + let qid = RelId (qualid_of_ident ?loc:pat.loc id) in + Name id, Some qid + in + let avoid = ids_of_pattern avoid pat in + let avoid = add_name avoid na in + (avoid, (na, pat, expand) :: bnd) + in + let (_, bnd) = List.fold_left fold (avoid, []) bnd in + let fold e (na, pat, expand) = match expand with + | None -> e + | Some qid -> + let loc = loc_of_relid qid in + CAst.make ?loc @@ CTacCse (CAst.make ?loc @@ CTacRef qid, [pat, e]) + in + let expand e = List.fold_left fold e bnd in + let nas = List.rev_map (fun (na, _, _) -> na) bnd in + (nas, expand) + +let is_alias env qid = match get_variable env qid with +| ArgArg (TacAlias _) -> true +| ArgVar _ | (ArgArg (TacConstant _)) -> false + +let rec intern_rec env {loc;v=e} = match e with +| CTacAtm atm -> intern_atm env atm +| CTacRef qid -> + begin match get_variable env qid with + | ArgVar {CAst.v=id} -> + let sch = Id.Map.find id env.env_var in + (GTacVar id, fresh_mix_type_scheme env sch) + | ArgArg (TacConstant kn) -> + let { Tac2env.gdata_type = sch } = + try Tac2env.interp_global kn + with Not_found -> + CErrors.anomaly (str "Missing hardwired primitive " ++ KerName.print kn) + in + (GTacRef kn, fresh_type_scheme env sch) + | ArgArg (TacAlias kn) -> + let e = + try Tac2env.interp_alias kn + with Not_found -> + CErrors.anomaly (str "Missing hardwired alias " ++ KerName.print kn) + in + intern_rec env e + end +| CTacCst qid -> + let kn = get_constructor env qid in + intern_constructor env loc kn [] +| CTacFun (bnd, e) -> + let bnd = List.map extract_pattern_type bnd in + let map (_, t) = match t with + | None -> GTypVar (fresh_id env) + | Some t -> intern_type env t + in + let tl = List.map map bnd in + let (nas, exp) = expand_pattern (Id.Map.domain env.env_var) bnd in + let env = List.fold_left2 (fun env na t -> push_name na (monomorphic t) env) env nas tl in + let (e, t) = intern_rec env (exp e) in + let t = List.fold_right (fun t accu -> GTypArrow (t, accu)) tl t in + (GTacFun (nas, e), t) +| CTacApp ({loc;v=CTacCst qid}, args) -> + let kn = get_constructor env qid in + intern_constructor env loc kn args +| CTacApp ({v=CTacRef qid}, args) when is_alias env qid -> + let kn = match get_variable env qid with + | ArgArg (TacAlias kn) -> kn + | ArgVar _ | (ArgArg (TacConstant _)) -> assert false + in + let e = Tac2env.interp_alias kn in + let map arg = + (* Thunk alias arguments *) + let loc = arg.loc in + let t_unit = CAst.make ?loc @@ CTypRef (AbsKn (Tuple 0), []) in + let var = CAst.make ?loc @@ CPatCnv (CAst.make ?loc @@ CPatVar Anonymous, t_unit) in + CAst.make ?loc @@ CTacFun ([var], arg) + in + let args = List.map map args in + intern_rec env (CAst.make ?loc @@ CTacApp (e, args)) +| CTacApp (f, args) -> + let loc = f.loc in + let (f, ft) = intern_rec env f in + let fold arg (args, t) = + let loc = arg.loc in + let (arg, argt) = intern_rec env arg in + (arg :: args, (loc, argt) :: t) + in + let (args, t) = List.fold_right fold args ([], []) in + let ret = unify_arrow ?loc env ft t in + (GTacApp (f, args), ret) +| CTacLet (is_rec, el, e) -> + let map (pat, e) = + let (pat, ty) = extract_pattern_type pat in + (pat, ty, e) + in + let el = List.map map el in + let fold accu (pat, _, e) = + let ids = ids_of_pattern Id.Set.empty pat in + let common = Id.Set.inter ids accu in + if Id.Set.is_empty common then Id.Set.union ids accu + else + let id = Id.Set.choose common in + user_err ?loc:pat.loc (str "Variable " ++ Id.print id ++ str " is bound several \ + times in this matching") + in + let ids = List.fold_left fold Id.Set.empty el in + if is_rec then intern_let_rec env loc ids el e + else intern_let env loc ids el e +| CTacCnv (e, tc) -> + let (e, t) = intern_rec env e in + let tc = intern_type env tc in + let () = unify ?loc env t tc in + (e, tc) +| CTacSeq (e1, e2) -> + let loc1 = e1.loc in + let (e1, t1) = intern_rec env e1 in + let (e2, t2) = intern_rec env e2 in + let () = check_elt_unit loc1 env t1 in + (GTacLet (false, [Anonymous, e1], e2), t2) +| CTacCse (e, pl) -> + intern_case env loc e pl +| CTacRec fs -> + intern_record env loc fs +| CTacPrj (e, proj) -> + let pinfo = get_projection proj in + let loc = e.loc in + let (e, t) = intern_rec env e in + let subst = Array.init pinfo.pdata_prms (fun _ -> fresh_id env) in + let params = Array.map_to_list (fun i -> GTypVar i) subst in + let exp = GTypRef (Other pinfo.pdata_type, params) in + let () = unify ?loc env t exp in + let substf i = GTypVar subst.(i) in + let ret = subst_type substf pinfo.pdata_ptyp in + (GTacPrj (pinfo.pdata_type, e, pinfo.pdata_indx), ret) +| CTacSet (e, proj, r) -> + let pinfo = get_projection proj in + let () = + if not pinfo.pdata_mutb then + let loc = match proj with + | RelId {CAst.loc} -> loc + | AbsKn _ -> None + in + user_err ?loc (str "Field is not mutable") + in + let subst = Array.init pinfo.pdata_prms (fun _ -> fresh_id env) in + let params = Array.map_to_list (fun i -> GTypVar i) subst in + let exp = GTypRef (Other pinfo.pdata_type, params) in + let e = intern_rec_with_constraint env e exp in + let substf i = GTypVar subst.(i) in + let ret = subst_type substf pinfo.pdata_ptyp in + let r = intern_rec_with_constraint env r ret in + (GTacSet (pinfo.pdata_type, e, pinfo.pdata_indx, r), GTypRef (Tuple 0, [])) +| CTacExt (tag, arg) -> + let open Genintern in + let self ist e = + let env = match Store.get ist.extra ltac2_env with + | None -> empty_env () + | Some env -> env + in + intern_rec env e + in + let obj = interp_ml_object tag in + (* External objects do not have access to the named context because this is + not stable by dynamic semantics. *) + let genv = Global.env_of_context Environ.empty_named_context_val in + let ist = empty_glob_sign genv in + let ist = { ist with extra = Store.set ist.extra ltac2_env env } in + let arg, tpe = + if env.env_str then + let arg () = obj.ml_intern self ist arg in + Flags.with_option Ltac_plugin.Tacintern.strict_check arg () + else + obj.ml_intern self ist arg + in + let e = match arg with + | GlbVal arg -> GTacExt (tag, arg) + | GlbTacexpr e -> e + in + (e, tpe) + +and intern_rec_with_constraint env e exp = + let (er, t) = intern_rec env e in + let () = unify ?loc:e.loc env t exp in + er + +and intern_let env loc ids el e = + let avoid = Id.Set.union ids (Id.Map.domain env.env_var) in + let fold (pat, t, e) (avoid, accu) = + let nas, exp = expand_pattern avoid [pat, t] in + let na = match nas with [x] -> x | _ -> assert false in + let avoid = List.fold_left add_name avoid nas in + (avoid, (na, exp, t, e) :: accu) + in + let (_, el) = List.fold_right fold el (avoid, []) in + let fold (na, exp, tc, e) (body, el, p) = + let (e, t) = match tc with + | None -> intern_rec env e + | Some tc -> + let tc = intern_type env tc in + (intern_rec_with_constraint env e tc, tc) + in + let t = if is_value e then abstract_var env t else monomorphic t in + (exp body, (na, e) :: el, (na, t) :: p) + in + let (e, el, p) = List.fold_right fold el (e, [], []) in + let env = List.fold_left (fun accu (na, t) -> push_name na t accu) env p in + let (e, t) = intern_rec env e in + (GTacLet (false, el, e), t) + +and intern_let_rec env loc ids el e = + let map env (pat, t, e) = + let na = match pat.v with + | CPatVar na -> na + | CPatRef _ | CPatCnv _ -> + user_err ?loc:pat.loc (str "This kind of pattern is forbidden in let-rec bindings") + in + let id = fresh_id env in + let env = push_name na (monomorphic (GTypVar id)) env in + (env, (loc, na, t, e, id)) + in + let (env, el) = List.fold_left_map map env el in + let fold (loc, na, tc, e, id) (el, tl) = + let loc_e = e.loc in + let (e, t) = intern_rec env e in + let () = + if not (is_rec_rhs e) then + user_err ?loc:loc_e (str "This kind of expression is not allowed as \ + right-hand side of a recursive binding") + in + let () = unify ?loc env t (GTypVar id) in + let () = match tc with + | None -> () + | Some tc -> + let tc = intern_type env tc in + unify ?loc env t tc + in + ((na, e) :: el, t :: tl) + in + let (el, tl) = List.fold_right fold el ([], []) in + let (e, t) = intern_rec env e in + (GTacLet (true, el, e), t) + +(** For now, patterns recognized by the pattern-matching compiling are limited + to depth-one where leaves are either variables or catch-all *) +and intern_case env loc e pl = + let (e', t) = intern_rec env e in + let todo ?loc () = user_err ?loc (str "Pattern not handled yet") in + match get_pattern_kind env pl with + | PKind_any -> + let (pat, b) = List.hd pl in + let na = match intern_patexpr env pat with + | GPatVar na -> na + | _ -> assert false + in + let () = check_redundant_clause (List.tl pl) in + let env = push_name na (monomorphic t) env in + let (b, tb) = intern_rec env b in + (GTacLet (false, [na, e'], b), tb) + | PKind_empty -> + let kn = check_elt_empty loc env t in + let r = fresh_id env in + (GTacCse (e', Other kn, [||], [||]), GTypVar r) + | PKind_variant kn -> + let subst, tc = fresh_reftype env kn in + let () = unify ?loc:e.loc env t tc in + let (nconst, nnonconst, arities) = match kn with + | Tuple 0 -> 1, 0, [0] + | Tuple n -> 0, 1, [n] + | Other kn -> + let (_, def) = Tac2env.interp_type kn in + let galg = match def with | GTydAlg c -> c | _ -> assert false in + let arities = List.map (fun (_, args) -> List.length args) galg.galg_constructors in + galg.galg_nconst, galg.galg_nnonconst, arities + in + let const = Array.make nconst None in + let nonconst = Array.make nnonconst None in + let ret = GTypVar (fresh_id env) in + let rec intern_branch = function + | [] -> () + | (pat, br) :: rem -> + let tbr = match pat.v with + | CPatVar (Name _) -> + let loc = pat.loc in + todo ?loc () + | CPatVar Anonymous -> + let () = check_redundant_clause rem in + let (br', brT) = intern_rec env br in + (* Fill all remaining branches *) + let fill (ncst, narg) arity = + if Int.equal arity 0 then + let () = + if Option.is_empty const.(ncst) then const.(ncst) <- Some br' + in + (succ ncst, narg) + else + let () = + if Option.is_empty nonconst.(narg) then + let ids = Array.make arity Anonymous in + nonconst.(narg) <- Some (ids, br') + in + (ncst, succ narg) + in + let _ = List.fold_left fill (0, 0) arities in + brT + | CPatRef (qid, args) -> + let loc = pat.loc in + let knc = get_constructor env qid in + let kn', index, arity = match knc with + | Tuple n -> Tuple n, 0, List.init n (fun i -> GTypVar i) + | Other knc -> + let data = Tac2env.interp_constructor knc in + let index = Option.get data.cdata_indx in + Other data.cdata_type, index, data.cdata_args + in + let () = + if not (eq_or_tuple KerName.equal kn kn') then + invalid_pattern ?loc kn kn' + in + let get_id pat = match pat with + | {v=CPatVar na} -> na + | {loc} -> todo ?loc () + in + let ids = List.map get_id args in + let nids = List.length ids in + let nargs = List.length arity in + let () = match knc with + | Tuple n -> assert (n == nids) + | Other knc -> + if not (Int.equal nids nargs) then error_nargs_mismatch ?loc knc nargs nids + in + let fold env id tpe = + (* Instantiate all arguments *) + let subst n = GTypVar subst.(n) in + let tpe = subst_type subst tpe in + push_name id (monomorphic tpe) env + in + let nenv = List.fold_left2 fold env ids arity in + let (br', brT) = intern_rec nenv br in + let () = + if List.is_empty args then + if Option.is_empty const.(index) then const.(index) <- Some br' + else warn_redundant_clause ?loc () + else + let ids = Array.of_list ids in + if Option.is_empty nonconst.(index) then nonconst.(index) <- Some (ids, br') + else warn_redundant_clause ?loc () + in + brT + | CPatCnv _ -> + user_err ?loc (str "Pattern not handled yet") + in + let () = unify ?loc:br.loc env tbr ret in + intern_branch rem + in + let () = intern_branch pl in + let map n is_const = function + | None -> + let kn = match kn with Other kn -> kn | _ -> assert false in + let cstr = pr_internal_constructor kn n is_const in + user_err ?loc (str "Unhandled match case for constructor " ++ cstr) + | Some x -> x + in + let const = Array.mapi (fun i o -> map i true o) const in + let nonconst = Array.mapi (fun i o -> map i false o) nonconst in + let ce = GTacCse (e', kn, const, nonconst) in + (ce, ret) + | PKind_open kn -> + let subst, tc = fresh_reftype env (Other kn) in + let () = unify ?loc:e.loc env t tc in + let ret = GTypVar (fresh_id env) in + let rec intern_branch map = function + | [] -> + user_err ?loc (str "Missing default case") + | (pat, br) :: rem -> + match intern_patexpr env pat with + | GPatVar na -> + let () = check_redundant_clause rem in + let nenv = push_name na (monomorphic tc) env in + let br' = intern_rec_with_constraint nenv br ret in + let def = (na, br') in + (map, def) + | GPatRef (knc, args) -> + let get = function + | GPatVar na -> na + | GPatRef _ -> + user_err ?loc (str "TODO: Unhandled match case") (* FIXME *) + in + let loc = pat.loc in + let knc = match knc with + | Other knc -> knc + | Tuple n -> invalid_pattern ?loc (Other kn) (Tuple n) + in + let ids = List.map get args in + let data = Tac2env.interp_constructor knc in + let () = + if not (KerName.equal kn data.cdata_type) then + invalid_pattern ?loc (Other kn) (Other data.cdata_type) + in + let nids = List.length ids in + let nargs = List.length data.cdata_args in + let () = + if not (Int.equal nids nargs) then error_nargs_mismatch ?loc knc nargs nids + in + let fold env id tpe = + (* Instantiate all arguments *) + let subst n = GTypVar subst.(n) in + let tpe = subst_type subst tpe in + push_name id (monomorphic tpe) env + in + let nenv = List.fold_left2 fold env ids data.cdata_args in + let br' = intern_rec_with_constraint nenv br ret in + let map = + if KNmap.mem knc map then + let () = warn_redundant_clause ?loc () in + map + else + KNmap.add knc (Anonymous, Array.of_list ids, br') map + in + intern_branch map rem + in + let (map, def) = intern_branch KNmap.empty pl in + (GTacWth { opn_match = e'; opn_branch = map; opn_default = def }, ret) + +and intern_constructor env loc kn args = match kn with +| Other kn -> + let cstr = interp_constructor kn in + let nargs = List.length cstr.cdata_args in + if Int.equal nargs (List.length args) then + let subst = Array.init cstr.cdata_prms (fun _ -> fresh_id env) in + let substf i = GTypVar subst.(i) in + let types = List.map (fun t -> subst_type substf t) cstr.cdata_args in + let targs = List.init cstr.cdata_prms (fun i -> GTypVar subst.(i)) in + let ans = GTypRef (Other cstr.cdata_type, targs) in + let map arg tpe = intern_rec_with_constraint env arg tpe in + let args = List.map2 map args types in + match cstr.cdata_indx with + | Some idx -> + (GTacCst (Other cstr.cdata_type, idx, args), ans) + | None -> + (GTacOpn (kn, args), ans) + else + error_nargs_mismatch ?loc kn nargs (List.length args) +| Tuple n -> + assert (Int.equal n (List.length args)); + let types = List.init n (fun i -> GTypVar (fresh_id env)) in + let map arg tpe = intern_rec_with_constraint env arg tpe in + let args = List.map2 map args types in + let ans = GTypRef (Tuple n, types) in + GTacCst (Tuple n, 0, args), ans + +and intern_record env loc fs = + let map (proj, e) = + let loc = match proj with + | RelId {CAst.loc} -> loc + | AbsKn _ -> None + in + let proj = get_projection proj in + (loc, proj, e) + in + let fs = List.map map fs in + let kn = match fs with + | [] -> user_err ?loc (str "Cannot infer the corresponding record type") + | (_, proj, _) :: _ -> proj.pdata_type + in + let params, typdef = match Tac2env.interp_type kn with + | n, GTydRec def -> n, def + | _ -> assert false + in + let subst = Array.init params (fun _ -> fresh_id env) in + (* Set the answer [args] imperatively *) + let args = Array.make (List.length typdef) None in + let iter (loc, pinfo, e) = + if KerName.equal kn pinfo.pdata_type then + let index = pinfo.pdata_indx in + match args.(index) with + | None -> + let exp = subst_type (fun i -> GTypVar subst.(i)) pinfo.pdata_ptyp in + let e = intern_rec_with_constraint env e exp in + args.(index) <- Some e + | Some _ -> + let (name, _, _) = List.nth typdef pinfo.pdata_indx in + user_err ?loc (str "Field " ++ Id.print name ++ str " is defined \ + several times") + else + user_err ?loc (str "Field " ++ (*KerName.print knp ++*) str " does not \ + pertain to record definition " ++ pr_typref pinfo.pdata_type) + in + let () = List.iter iter fs in + let () = match Array.findi (fun _ o -> Option.is_empty o) args with + | None -> () + | Some i -> + let (field, _, _) = List.nth typdef i in + user_err ?loc (str "Field " ++ Id.print field ++ str " is undefined") + in + let args = Array.map_to_list Option.get args in + let tparam = List.init params (fun i -> GTypVar subst.(i)) in + (GTacCst (Other kn, 0, args), GTypRef (Other kn, tparam)) + +let normalize env (count, vars) (t : UF.elt glb_typexpr) = + let get_var id = + try UF.Map.find id !vars + with Not_found -> + let () = assert env.env_opn in + let n = GTypVar !count in + let () = incr count in + let () = vars := UF.Map.add id n !vars in + n + in + let rec subst id = match UF.find id env.env_cst with + | id, None -> get_var id + | _, Some t -> subst_type subst t + in + subst_type subst t + +let intern ~strict e = + let env = empty_env () in + let env = if strict then env else { env with env_str = false } in + let (e, t) = intern_rec env e in + let count = ref 0 in + let vars = ref UF.Map.empty in + let t = normalize env (count, vars) t in + (e, (!count, t)) + +let intern_typedef self (ids, t) : glb_quant_typedef = + let env = { (empty_env ()) with env_rec = self } in + (* Initialize type parameters *) + let map id = get_alias id env in + let ids = List.map map ids in + let count = ref (List.length ids) in + let vars = ref UF.Map.empty in + let iter n id = vars := UF.Map.add id (GTypVar n) !vars in + let () = List.iteri iter ids in + (* Do not accept unbound type variables *) + let env = { env with env_opn = false } in + let intern t = + let t = intern_type env t in + normalize env (count, vars) t + in + let count = !count in + match t with + | CTydDef None -> (count, GTydDef None) + | CTydDef (Some t) -> (count, GTydDef (Some (intern t))) + | CTydAlg constrs -> + let map (c, t) = (c, List.map intern t) in + let constrs = List.map map constrs in + let getn (const, nonconst) (c, args) = match args with + | [] -> (succ const, nonconst) + | _ :: _ -> (const, succ nonconst) + in + let nconst, nnonconst = List.fold_left getn (0, 0) constrs in + let galg = { + galg_constructors = constrs; + galg_nconst = nconst; + galg_nnonconst = nnonconst; + } in + (count, GTydAlg galg) + | CTydRec fields -> + let map (c, mut, t) = (c, mut, intern t) in + let fields = List.map map fields in + (count, GTydRec fields) + | CTydOpn -> (count, GTydOpn) + +let intern_open_type t = + let env = empty_env () in + let t = intern_type env t in + let count = ref 0 in + let vars = ref UF.Map.empty in + let t = normalize env (count, vars) t in + (!count, t) + +(** Subtyping *) + +let check_subtype t1 t2 = + let env = empty_env () in + let t1 = fresh_type_scheme env t1 in + (* We build a substitution mimicking rigid variable by using dummy tuples *) + let rigid i = GTypRef (Tuple (i + 1), []) in + let (n, t2) = t2 in + let subst = Array.init n rigid in + let substf i = subst.(i) in + let t2 = subst_type substf t2 in + try unify0 env t1 t2; true with CannotUnify _ -> false + +(** Globalization *) + +let get_projection0 var = match var with +| RelId qid -> + let kn = try Tac2env.locate_projection qid with Not_found -> + user_err ?loc:qid.CAst.loc (pr_qualid qid ++ str " is not a projection") + in + kn +| AbsKn kn -> kn + +let rec globalize ids ({loc;v=er} as e) = match er with +| CTacAtm _ -> e +| CTacRef ref -> + let mem id = Id.Set.mem id ids in + begin match get_variable0 mem ref with + | ArgVar _ -> e + | ArgArg kn -> CAst.make ?loc @@ CTacRef (AbsKn kn) + end +| CTacCst qid -> + let knc = get_constructor () qid in + CAst.make ?loc @@ CTacCst (AbsKn knc) +| CTacFun (bnd, e) -> + let fold (pats, accu) pat = + let accu = ids_of_pattern accu pat in + let pat = globalize_pattern ids pat in + (pat :: pats, accu) + in + let bnd, ids = List.fold_left fold ([], ids) bnd in + let bnd = List.rev bnd in + let e = globalize ids e in + CAst.make ?loc @@ CTacFun (bnd, e) +| CTacApp (e, el) -> + let e = globalize ids e in + let el = List.map (fun e -> globalize ids e) el in + CAst.make ?loc @@ CTacApp (e, el) +| CTacLet (isrec, bnd, e) -> + let fold accu (pat, _) = ids_of_pattern accu pat in + let ext = List.fold_left fold Id.Set.empty bnd in + let eids = Id.Set.union ext ids in + let e = globalize eids e in + let map (qid, e) = + let ids = if isrec then eids else ids in + let qid = globalize_pattern ids qid in + (qid, globalize ids e) + in + let bnd = List.map map bnd in + CAst.make ?loc @@ CTacLet (isrec, bnd, e) +| CTacCnv (e, t) -> + let e = globalize ids e in + CAst.make ?loc @@ CTacCnv (e, t) +| CTacSeq (e1, e2) -> + let e1 = globalize ids e1 in + let e2 = globalize ids e2 in + CAst.make ?loc @@ CTacSeq (e1, e2) +| CTacCse (e, bl) -> + let e = globalize ids e in + let bl = List.map (fun b -> globalize_case ids b) bl in + CAst.make ?loc @@ CTacCse (e, bl) +| CTacRec r -> + let map (p, e) = + let p = get_projection0 p in + let e = globalize ids e in + (AbsKn p, e) + in + CAst.make ?loc @@ CTacRec (List.map map r) +| CTacPrj (e, p) -> + let e = globalize ids e in + let p = get_projection0 p in + CAst.make ?loc @@ CTacPrj (e, AbsKn p) +| CTacSet (e, p, e') -> + let e = globalize ids e in + let p = get_projection0 p in + let e' = globalize ids e' in + CAst.make ?loc @@ CTacSet (e, AbsKn p, e') +| CTacExt (tag, arg) -> + let arg = str (Tac2dyn.Arg.repr tag) in + CErrors.user_err ?loc (str "Cannot globalize generic arguments of type" ++ spc () ++ arg) + +and globalize_case ids (p, e) = + (globalize_pattern ids p, globalize ids e) + +and globalize_pattern ids ({loc;v=pr} as p) = match pr with +| CPatVar _ -> p +| CPatRef (cst, pl) -> + let knc = get_constructor () cst in + let cst = AbsKn knc in + let pl = List.map (fun p -> globalize_pattern ids p) pl in + CAst.make ?loc @@ CPatRef (cst, pl) +| CPatCnv (pat, ty) -> + let pat = globalize_pattern ids pat in + CAst.make ?loc @@ CPatCnv (pat, ty) + +(** Kernel substitution *) + +open Mod_subst + +let subst_or_tuple f subst o = match o with +| Tuple _ -> o +| Other v -> + let v' = f subst v in + if v' == v then o else Other v' + +let rec subst_type subst t = match t with +| GTypVar _ -> t +| GTypArrow (t1, t2) -> + let t1' = subst_type subst t1 in + let t2' = subst_type subst t2 in + if t1' == t1 && t2' == t2 then t + else GTypArrow (t1', t2') +| GTypRef (kn, tl) -> + let kn' = subst_or_tuple subst_kn subst kn in + let tl' = List.Smart.map (fun t -> subst_type subst t) tl in + if kn' == kn && tl' == tl then t else GTypRef (kn', tl') + +let rec subst_expr subst e = match e with +| GTacAtm _ | GTacVar _ | GTacPrm _ -> e +| GTacRef kn -> GTacRef (subst_kn subst kn) +| GTacFun (ids, e) -> GTacFun (ids, subst_expr subst e) +| GTacApp (f, args) -> + GTacApp (subst_expr subst f, List.map (fun e -> subst_expr subst e) args) +| GTacLet (r, bs, e) -> + let bs = List.map (fun (na, e) -> (na, subst_expr subst e)) bs in + GTacLet (r, bs, subst_expr subst e) +| GTacCst (t, n, el) as e0 -> + let t' = subst_or_tuple subst_kn subst t in + let el' = List.Smart.map (fun e -> subst_expr subst e) el in + if t' == t && el' == el then e0 else GTacCst (t', n, el') +| GTacCse (e, ci, cse0, cse1) -> + let cse0' = Array.map (fun e -> subst_expr subst e) cse0 in + let cse1' = Array.map (fun (ids, e) -> (ids, subst_expr subst e)) cse1 in + let ci' = subst_or_tuple subst_kn subst ci in + GTacCse (subst_expr subst e, ci', cse0', cse1') +| GTacWth { opn_match = e; opn_branch = br; opn_default = (na, def) } as e0 -> + let e' = subst_expr subst e in + let def' = subst_expr subst def in + let fold kn (self, vars, p) accu = + let kn' = subst_kn subst kn in + let p' = subst_expr subst p in + if kn' == kn && p' == p then accu + else KNmap.add kn' (self, vars, p') (KNmap.remove kn accu) + in + let br' = KNmap.fold fold br br in + if e' == e && br' == br && def' == def then e0 + else GTacWth { opn_match = e'; opn_default = (na, def'); opn_branch = br' } +| GTacPrj (kn, e, p) as e0 -> + let kn' = subst_kn subst kn in + let e' = subst_expr subst e in + if kn' == kn && e' == e then e0 else GTacPrj (kn', e', p) +| GTacSet (kn, e, p, r) as e0 -> + let kn' = subst_kn subst kn in + let e' = subst_expr subst e in + let r' = subst_expr subst r in + if kn' == kn && e' == e && r' == r then e0 else GTacSet (kn', e', p, r') +| GTacExt (tag, arg) -> + let tpe = interp_ml_object tag in + let arg' = tpe.ml_subst subst arg in + if arg' == arg then e else GTacExt (tag, arg') +| GTacOpn (kn, el) as e0 -> + let kn' = subst_kn subst kn in + let el' = List.Smart.map (fun e -> subst_expr subst e) el in + if kn' == kn && el' == el then e0 else GTacOpn (kn', el') + +let subst_typedef subst e = match e with +| GTydDef t -> + let t' = Option.Smart.map (fun t -> subst_type subst t) t in + if t' == t then e else GTydDef t' +| GTydAlg galg -> + let map (c, tl as p) = + let tl' = List.Smart.map (fun t -> subst_type subst t) tl in + if tl' == tl then p else (c, tl') + in + let constrs' = List.Smart.map map galg.galg_constructors in + if constrs' == galg.galg_constructors then e + else GTydAlg { galg with galg_constructors = constrs' } +| GTydRec fields -> + let map (c, mut, t as p) = + let t' = subst_type subst t in + if t' == t then p else (c, mut, t') + in + let fields' = List.Smart.map map fields in + if fields' == fields then e else GTydRec fields' +| GTydOpn -> GTydOpn + +let subst_quant_typedef subst (prm, def as qdef) = + let def' = subst_typedef subst def in + if def' == def then qdef else (prm, def') + +let subst_type_scheme subst (prm, t as sch) = + let t' = subst_type subst t in + if t' == t then sch else (prm, t') + +let subst_or_relid subst ref = match ref with +| RelId _ -> ref +| AbsKn kn -> + let kn' = subst_or_tuple subst_kn subst kn in + if kn' == kn then ref else AbsKn kn' + +let rec subst_rawtype subst ({loc;v=tr} as t) = match tr with +| CTypVar _ -> t +| CTypArrow (t1, t2) -> + let t1' = subst_rawtype subst t1 in + let t2' = subst_rawtype subst t2 in + if t1' == t1 && t2' == t2 then t else CAst.make ?loc @@ CTypArrow (t1', t2') +| CTypRef (ref, tl) -> + let ref' = subst_or_relid subst ref in + let tl' = List.Smart.map (fun t -> subst_rawtype subst t) tl in + if ref' == ref && tl' == tl then t else CAst.make ?loc @@ CTypRef (ref', tl') + +let subst_tacref subst ref = match ref with +| RelId _ -> ref +| AbsKn (TacConstant kn) -> + let kn' = subst_kn subst kn in + if kn' == kn then ref else AbsKn (TacConstant kn') +| AbsKn (TacAlias kn) -> + let kn' = subst_kn subst kn in + if kn' == kn then ref else AbsKn (TacAlias kn') + +let subst_projection subst prj = match prj with +| RelId _ -> prj +| AbsKn kn -> + let kn' = subst_kn subst kn in + if kn' == kn then prj else AbsKn kn' + +let rec subst_rawpattern subst ({loc;v=pr} as p) = match pr with +| CPatVar _ -> p +| CPatRef (c, pl) -> + let pl' = List.Smart.map (fun p -> subst_rawpattern subst p) pl in + let c' = subst_or_relid subst c in + if pl' == pl && c' == c then p else CAst.make ?loc @@ CPatRef (c', pl') +| CPatCnv (pat, ty) -> + let pat' = subst_rawpattern subst pat in + let ty' = subst_rawtype subst ty in + if pat' == pat && ty' == ty then p else CAst.make ?loc @@ CPatCnv (pat', ty') + +(** Used for notations *) +let rec subst_rawexpr subst ({loc;v=tr} as t) = match tr with +| CTacAtm _ -> t +| CTacRef ref -> + let ref' = subst_tacref subst ref in + if ref' == ref then t else CAst.make ?loc @@ CTacRef ref' +| CTacCst ref -> + let ref' = subst_or_relid subst ref in + if ref' == ref then t else CAst.make ?loc @@ CTacCst ref' +| CTacFun (bnd, e) -> + let map pat = subst_rawpattern subst pat in + let bnd' = List.Smart.map map bnd in + let e' = subst_rawexpr subst e in + if bnd' == bnd && e' == e then t else CAst.make ?loc @@ CTacFun (bnd', e') +| CTacApp (e, el) -> + let e' = subst_rawexpr subst e in + let el' = List.Smart.map (fun e -> subst_rawexpr subst e) el in + if e' == e && el' == el then t else CAst.make ?loc @@ CTacApp (e', el') +| CTacLet (isrec, bnd, e) -> + let map (na, e as p) = + let na' = subst_rawpattern subst na in + let e' = subst_rawexpr subst e in + if na' == na && e' == e then p else (na', e') + in + let bnd' = List.Smart.map map bnd in + let e' = subst_rawexpr subst e in + if bnd' == bnd && e' == e then t else CAst.make ?loc @@ CTacLet (isrec, bnd', e') +| CTacCnv (e, c) -> + let e' = subst_rawexpr subst e in + let c' = subst_rawtype subst c in + if c' == c && e' == e then t else CAst.make ?loc @@ CTacCnv (e', c') +| CTacSeq (e1, e2) -> + let e1' = subst_rawexpr subst e1 in + let e2' = subst_rawexpr subst e2 in + if e1' == e1 && e2' == e2 then t else CAst.make ?loc @@ CTacSeq (e1', e2') +| CTacCse (e, bl) -> + let map (p, e as x) = + let p' = subst_rawpattern subst p in + let e' = subst_rawexpr subst e in + if p' == p && e' == e then x else (p', e') + in + let e' = subst_rawexpr subst e in + let bl' = List.Smart.map map bl in + if e' == e && bl' == bl then t else CAst.make ?loc @@ CTacCse (e', bl') +| CTacRec el -> + let map (prj, e as p) = + let prj' = subst_projection subst prj in + let e' = subst_rawexpr subst e in + if prj' == prj && e' == e then p else (prj', e') + in + let el' = List.Smart.map map el in + if el' == el then t else CAst.make ?loc @@ CTacRec el' +| CTacPrj (e, prj) -> + let prj' = subst_projection subst prj in + let e' = subst_rawexpr subst e in + if prj' == prj && e' == e then t else CAst.make ?loc @@ CTacPrj (e', prj') +| CTacSet (e, prj, r) -> + let prj' = subst_projection subst prj in + let e' = subst_rawexpr subst e in + let r' = subst_rawexpr subst r in + if prj' == prj && e' == e && r' == r then t else CAst.make ?loc @@ CTacSet (e', prj', r') +| CTacExt _ -> assert false (** Should not be generated by globalization *) + +(** Registering *) + +let () = + let open Genintern in + let intern ist tac = + let env = match Genintern.Store.get ist.extra ltac2_env with + | None -> + (* Only happens when Ltac2 is called from a constr or ltac1 quotation *) + let env = empty_env () in + if !Ltac_plugin.Tacintern.strict_check then env + else { env with env_str = false } + | Some env -> env + in + let loc = tac.loc in + let (tac, t) = intern_rec env tac in + let () = check_elt_unit loc env t in + (ist, tac) + in + Genintern.register_intern0 wit_ltac2 intern +let () = Genintern.register_subst0 wit_ltac2 subst_expr + +let () = + let open Genintern in + let intern ist (loc, id) = + let env = match Genintern.Store.get ist.extra ltac2_env with + | None -> + (* Only happens when Ltac2 is called from a constr or ltac1 quotation *) + let env = empty_env () in + if !Ltac_plugin.Tacintern.strict_check then env + else { env with env_str = false } + | Some env -> env + in + let t = + try Id.Map.find id env.env_var + with Not_found -> + CErrors.user_err ?loc (str "Unbound value " ++ Id.print id) + in + let t = fresh_mix_type_scheme env t in + let () = unify ?loc env t (GTypRef (Other t_constr, [])) in + (ist, id) + in + Genintern.register_intern0 wit_ltac2_quotation intern + +let () = Genintern.register_subst0 wit_ltac2_quotation (fun _ id -> id) diff --git a/vendor/Ltac2/src/tac2intern.mli b/vendor/Ltac2/src/tac2intern.mli new file mode 100644 index 0000000000..d646b5cda5 --- /dev/null +++ b/vendor/Ltac2/src/tac2intern.mli @@ -0,0 +1,46 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* raw_tacexpr -> glb_tacexpr * type_scheme +val intern_typedef : (KerName.t * int) Id.Map.t -> raw_quant_typedef -> glb_quant_typedef +val intern_open_type : raw_typexpr -> type_scheme + +(** Check that a term is a value. Only values are safe to marshall between + processes. *) +val is_value : glb_tacexpr -> bool +val check_unit : ?loc:Loc.t -> type_scheme -> unit + +val check_subtype : type_scheme -> type_scheme -> bool +(** [check_subtype t1 t2] returns [true] iff all values of intances of type [t1] + also have type [t2]. *) + +val subst_type : substitution -> 'a glb_typexpr -> 'a glb_typexpr +val subst_expr : substitution -> glb_tacexpr -> glb_tacexpr +val subst_quant_typedef : substitution -> glb_quant_typedef -> glb_quant_typedef +val subst_type_scheme : substitution -> type_scheme -> type_scheme + +val subst_rawexpr : substitution -> raw_tacexpr -> raw_tacexpr + +(** {5 Notations} *) + +val globalize : Id.Set.t -> raw_tacexpr -> raw_tacexpr +(** Replaces all qualified identifiers by their corresponding kernel name. The + set represents bound variables in the context. *) + +(** Errors *) + +val error_nargs_mismatch : ?loc:Loc.t -> ltac_constructor -> int -> int -> 'a +val error_nparams_mismatch : ?loc:Loc.t -> int -> int -> 'a + +(** Misc *) + +val drop_ltac2_env : Genintern.Store.t -> Genintern.Store.t diff --git a/vendor/Ltac2/src/tac2interp.ml b/vendor/Ltac2/src/tac2interp.ml new file mode 100644 index 0000000000..b0f8083aeb --- /dev/null +++ b/vendor/Ltac2/src/tac2interp.ml @@ -0,0 +1,227 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* >= fun sigma -> + match Evd.Store.get (Evd.get_extra_data sigma) backtrace with + | None -> Proofview.tclUNIT [] + | Some bt -> Proofview.tclUNIT bt + +let set_backtrace bt = + Proofview.tclEVARMAP >>= fun sigma -> + let store = Evd.get_extra_data sigma in + let store = Evd.Store.set store backtrace bt in + let sigma = Evd.set_extra_data store sigma in + Proofview.Unsafe.tclEVARS sigma + +let with_frame frame tac = + if !print_ltac2_backtrace then + get_backtrace >>= fun bt -> + set_backtrace (frame :: bt) >>= fun () -> + tac >>= fun ans -> + set_backtrace bt >>= fun () -> + Proofview.tclUNIT ans + else tac + +type environment = Tac2env.environment = { + env_ist : valexpr Id.Map.t; +} + +let empty_environment = { + env_ist = Id.Map.empty; +} + +type closure = { + mutable clos_env : valexpr Id.Map.t; + (** Mutable so that we can implement recursive functions imperatively *) + clos_var : Name.t list; + (** Bound variables *) + clos_exp : glb_tacexpr; + (** Body *) + clos_ref : ltac_constant option; + (** Global constant from which the closure originates *) +} + +let push_name ist id v = match id with +| Anonymous -> ist +| Name id -> { env_ist = Id.Map.add id v ist.env_ist } + +let get_var ist id = + try Id.Map.find id ist.env_ist with Not_found -> + anomaly (str "Unbound variable " ++ Id.print id) + +let get_ref ist kn = + try + let data = Tac2env.interp_global kn in + data.Tac2env.gdata_expr + with Not_found -> + anomaly (str "Unbound reference" ++ KerName.print kn) + +let return = Proofview.tclUNIT + +let rec interp (ist : environment) = function +| GTacAtm (AtmInt n) -> return (Tac2ffi.of_int n) +| GTacAtm (AtmStr s) -> return (Tac2ffi.of_string (Bytes.of_string s)) +| GTacVar id -> return (get_var ist id) +| GTacRef kn -> + let data = get_ref ist kn in + return (eval_pure (Some kn) data) +| GTacFun (ids, e) -> + let cls = { clos_ref = None; clos_env = ist.env_ist; clos_var = ids; clos_exp = e } in + let f = interp_app cls in + return (Tac2ffi.of_closure f) +| GTacApp (f, args) -> + interp ist f >>= fun f -> + Proofview.Monad.List.map (fun e -> interp ist e) args >>= fun args -> + Tac2ffi.apply (Tac2ffi.to_closure f) args +| GTacLet (false, el, e) -> + let fold accu (na, e) = + interp ist e >>= fun e -> + return (push_name accu na e) + in + Proofview.Monad.List.fold_left fold ist el >>= fun ist -> + interp ist e +| GTacLet (true, el, e) -> + let map (na, e) = match e with + | GTacFun (ids, e) -> + let cls = { clos_ref = None; clos_env = ist.env_ist; clos_var = ids; clos_exp = e } in + let f = Tac2ffi.of_closure (interp_app cls) in + na, cls, f + | _ -> anomaly (str "Ill-formed recursive function") + in + let fixs = List.map map el in + let fold accu (na, _, cls) = match na with + | Anonymous -> accu + | Name id -> { env_ist = Id.Map.add id cls accu.env_ist } + in + let ist = List.fold_left fold ist fixs in + (* Hack to make a cycle imperatively in the environment *) + let iter (_, e, _) = e.clos_env <- ist.env_ist in + let () = List.iter iter fixs in + interp ist e +| GTacCst (_, n, []) -> return (Valexpr.make_int n) +| GTacCst (_, n, el) -> + Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> + return (Valexpr.make_block n (Array.of_list el)) +| GTacCse (e, _, cse0, cse1) -> + interp ist e >>= fun e -> interp_case ist e cse0 cse1 +| GTacWth { opn_match = e; opn_branch = cse; opn_default = def } -> + interp ist e >>= fun e -> interp_with ist e cse def +| GTacPrj (_, e, p) -> + interp ist e >>= fun e -> interp_proj ist e p +| GTacSet (_, e, p, r) -> + interp ist e >>= fun e -> + interp ist r >>= fun r -> + interp_set ist e p r +| GTacOpn (kn, el) -> + Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> + return (Tac2ffi.of_open (kn, Array.of_list el)) +| GTacPrm (ml, el) -> + Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> + with_frame (FrPrim ml) (Tac2ffi.apply (Tac2env.interp_primitive ml) el) +| GTacExt (tag, e) -> + let tpe = Tac2env.interp_ml_object tag in + with_frame (FrExtn (tag, e)) (tpe.Tac2env.ml_interp ist e) + +and interp_app f = + let ans = fun args -> + let { clos_env = ist; clos_var = ids; clos_exp = e; clos_ref = kn } = f in + let frame = match kn with + | None -> FrAnon e + | Some kn -> FrLtac kn + in + let ist = { env_ist = ist } in + let ist = List.fold_left2 push_name ist ids args in + with_frame frame (interp ist e) + in + Tac2ffi.abstract (List.length f.clos_var) ans + +and interp_case ist e cse0 cse1 = + if Valexpr.is_int e then + interp ist cse0.(Tac2ffi.to_int e) + else + let (n, args) = Tac2ffi.to_block e in + let (ids, e) = cse1.(n) in + let ist = CArray.fold_left2 push_name ist ids args in + interp ist e + +and interp_with ist e cse def = + let (kn, args) = Tac2ffi.to_open e in + let br = try Some (KNmap.find kn cse) with Not_found -> None in + begin match br with + | None -> + let (self, def) = def in + let ist = push_name ist self e in + interp ist def + | Some (self, ids, p) -> + let ist = push_name ist self e in + let ist = CArray.fold_left2 push_name ist ids args in + interp ist p + end + +and interp_proj ist e p = + return (Valexpr.field e p) + +and interp_set ist e p r = + let () = Valexpr.set_field e p r in + return (Valexpr.make_int 0) + +and eval_pure kn = function +| GTacAtm (AtmInt n) -> Valexpr.make_int n +| GTacRef kn -> + let { Tac2env.gdata_expr = e } = + try Tac2env.interp_global kn + with Not_found -> assert false + in + eval_pure (Some kn) e +| GTacFun (na, e) -> + let cls = { clos_ref = kn; clos_env = Id.Map.empty; clos_var = na; clos_exp = e } in + let f = interp_app cls in + Tac2ffi.of_closure f +| GTacCst (_, n, []) -> Valexpr.make_int n +| GTacCst (_, n, el) -> Valexpr.make_block n (Array.map_of_list eval_unnamed el) +| GTacOpn (kn, el) -> Tac2ffi.of_open (kn, Array.map_of_list eval_unnamed el) +| GTacAtm (AtmStr _) | GTacLet _ | GTacVar _ | GTacSet _ +| GTacApp _ | GTacCse _ | GTacPrj _ | GTacPrm _ | GTacExt _ | GTacWth _ -> + anomaly (Pp.str "Term is not a syntactical value") + +and eval_unnamed e = eval_pure None e + + +(** Cross-boundary hacks. *) + +open Geninterp + +let val_env : environment Val.typ = Val.create "ltac2:env" +let env_ref = Id.of_string_soft "@@ltac2_env@@" + +let extract_env (Val.Dyn (tag, v)) : environment = +match Val.eq tag val_env with +| None -> assert false +| Some Refl -> v + +let get_env ist = + try extract_env (Id.Map.find env_ref ist) + with Not_found -> empty_environment + +let set_env env ist = + Id.Map.add env_ref (Val.Dyn (val_env, env)) ist diff --git a/vendor/Ltac2/src/tac2interp.mli b/vendor/Ltac2/src/tac2interp.mli new file mode 100644 index 0000000000..21fdcd03af --- /dev/null +++ b/vendor/Ltac2/src/tac2interp.mli @@ -0,0 +1,37 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* glb_tacexpr -> valexpr Proofview.tactic + +(* val interp_app : closure -> ml_tactic *) + +(** {5 Cross-boundary encodings} *) + +val get_env : Ltac_pretype.unbound_ltac_var_map -> environment +val set_env : environment -> Ltac_pretype.unbound_ltac_var_map -> Ltac_pretype.unbound_ltac_var_map + +(** {5 Exceptions} *) + +exception LtacError of KerName.t * valexpr array +(** Ltac2-defined exceptions seen from OCaml side *) + +(** {5 Backtrace} *) + +val get_backtrace : backtrace Proofview.tactic + +val with_frame : frame -> 'a Proofview.tactic -> 'a Proofview.tactic + +val print_ltac2_backtrace : bool ref diff --git a/vendor/Ltac2/src/tac2match.ml b/vendor/Ltac2/src/tac2match.ml new file mode 100644 index 0000000000..c9e549d47e --- /dev/null +++ b/vendor/Ltac2/src/tac2match.ml @@ -0,0 +1,232 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* None + | None, Some c | Some c, None -> Some c + | Some c1, Some c2 -> + if equal_instances env sigma c1 c2 then Some c1 + else raise Not_coherent_metas + in + Id.Map.merge merge s1 s2 + +let matching_error = + CErrors.UserError (Some "tactic matching" , Pp.str "No matching clauses for match.") + +let imatching_error = (matching_error, Exninfo.null) + +(** A functor is introduced to share the environment and the + evar_map. They do not change and it would be a pity to introduce + closures everywhere just for the occasional calls to + {!equal_instances}. *) +module type StaticEnvironment = sig + val env : Environ.env + val sigma : Evd.evar_map +end +module PatternMatching (E:StaticEnvironment) = struct + + + (** {6 The pattern-matching monad } *) + + + (** To focus on the algorithmic portion of pattern-matching, the + bookkeeping is relegated to a monad: the composition of the + bactracking monad of {!IStream.t} with a "writer" effect. *) + (* spiwack: as we don't benefit from the various stream optimisations + of Haskell, it may be costly to give the monad in direct style such as + here. We may want to use some continuation passing style. *) + type 'a tac = 'a Proofview.tactic + type 'a m = { stream : 'r. ('a -> result -> 'r tac) -> result -> 'r tac } + + (** The empty substitution. *) + let empty_subst = Id.Map.empty + + (** Composes two substitutions using {!verify_metas_coherence}. It + must be a monoid with neutral element {!empty_subst}. Raises + [Not_coherent_metas] when composition cannot be achieved. *) + let subst_prod s1 s2 = + if is_empty_subst s1 then s2 + else if is_empty_subst s2 then s1 + else verify_metas_coherence E.env E.sigma s1 s2 + + (** Merge two writers (and ignore the first value component). *) + let merge m1 m2 = + try Some { + subst = subst_prod m1.subst m2.subst; + } + with Not_coherent_metas -> None + + (** Monadic [return]: returns a single success with empty substitutions. *) + let return (type a) (lhs:a) : a m = + { stream = fun k ctx -> k lhs ctx } + + (** Monadic bind: each success of [x] is replaced by the successes + of [f x]. The substitutions of [x] and [f x] are composed, + dropping the apparent successes when the substitutions are not + coherent. *) + let (>>=) (type a) (type b) (m:a m) (f:a -> b m) : b m = + { stream = fun k ctx -> m.stream (fun x ctx -> (f x).stream k ctx) ctx } + + (** A variant of [(>>=)] when the first argument returns [unit]. *) + let (<*>) (type a) (m:unit m) (y:a m) : a m = + { stream = fun k ctx -> m.stream (fun () ctx -> y.stream k ctx) ctx } + + (** Failure of the pattern-matching monad: no success. *) + let fail (type a) : a m = { stream = fun _ _ -> Proofview.tclZERO matching_error } + + let run (m : 'a m) = + let ctx = { + subst = empty_subst ; + } in + let eval x ctx = Proofview.tclUNIT (x, ctx) in + m.stream eval ctx + + (** Chooses in a list, in the same order as the list *) + let rec pick (l:'a list) (e, info) : 'a m = match l with + | [] -> { stream = fun _ _ -> Proofview.tclZERO ~info e } + | x :: l -> + { stream = fun k ctx -> Proofview.tclOR (k x ctx) (fun e -> (pick l e).stream k ctx) } + + let pick l = pick l imatching_error + + let put_subst subst : unit m = + let s = { subst } in + { stream = fun k ctx -> match merge s ctx with None -> Proofview.tclZERO matching_error | Some s -> k () s } + + (** {6 Pattern-matching} *) + + let pattern_match_term pat term = + match pat with + | MatchPattern p -> + begin + try + put_subst (Constr_matching.matches E.env E.sigma p term) <*> + return None + with Constr_matching.PatternMatchingFailure -> fail + end + | MatchContext p -> + + let rec map s (e, info) = + { stream = fun k ctx -> match IStream.peek s with + | IStream.Nil -> Proofview.tclZERO ~info e + | IStream.Cons ({ Constr_matching.m_sub = (_, subst); m_ctx }, s) -> + let nctx = { subst } in + match merge ctx nctx with + | None -> (map s (e, info)).stream k ctx + | Some nctx -> Proofview.tclOR (k (Some (Lazy.force m_ctx)) nctx) (fun e -> (map s e).stream k ctx) + } + in + map (Constr_matching.match_subterm E.env E.sigma (Id.Set.empty,p) term) imatching_error + + let hyp_match_type pat hyps = + pick hyps >>= fun decl -> + let id = NamedDecl.get_id decl in + pattern_match_term pat (NamedDecl.get_type decl) >>= fun ctx -> + return (id, ctx) + + let _hyp_match_body_and_type bodypat typepat hyps = + pick hyps >>= function + | LocalDef (id,body,hyp) -> + pattern_match_term bodypat body >>= fun ctx_body -> + pattern_match_term typepat hyp >>= fun ctx_typ -> + return (id, ctx_body, ctx_typ) + | LocalAssum (id,hyp) -> fail + + let hyp_match pat hyps = + match pat with + | typepat -> + hyp_match_type typepat hyps +(* | Def ((_,hypname),bodypat,typepat) -> *) +(* hyp_match_body_and_type hypname bodypat typepat hyps *) + + (** [hyp_pattern_list_match pats hyps lhs], matches the list of + patterns [pats] against the hypotheses in [hyps], and eventually + returns [lhs]. *) + let rec hyp_pattern_list_match pats hyps accu = + match pats with + | pat::pats -> + hyp_match pat hyps >>= fun (matched_hyp, hyp_ctx) -> + let select_matched_hyp decl = Id.equal (NamedDecl.get_id decl) matched_hyp in + let hyps = CList.remove_first select_matched_hyp hyps in + hyp_pattern_list_match pats hyps ((matched_hyp, hyp_ctx) :: accu) + | [] -> return accu + + let rule_match_goal hyps concl = function + | (hyppats,conclpat) -> + (* the rules are applied from the topmost one (in the concrete + syntax) to the bottommost. *) + let hyppats = List.rev hyppats in + pattern_match_term conclpat concl >>= fun ctx_concl -> + hyp_pattern_list_match hyppats hyps [] >>= fun hyps -> + return (hyps, ctx_concl) + +end + +let match_goal env sigma concl ~rev rule = + let open Proofview.Notations in + let hyps = EConstr.named_context env in + let hyps = if rev then List.rev hyps else hyps in + let module E = struct + let env = env + let sigma = sigma + end in + let module M = PatternMatching(E) in + M.run (M.rule_match_goal hyps concl rule) >>= fun ((hyps, ctx_concl), subst) -> + Proofview.tclUNIT (hyps, ctx_concl, subst.subst) diff --git a/vendor/Ltac2/src/tac2match.mli b/vendor/Ltac2/src/tac2match.mli new file mode 100644 index 0000000000..c82c40d238 --- /dev/null +++ b/vendor/Ltac2/src/tac2match.mli @@ -0,0 +1,33 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* + Evd.evar_map -> + constr -> + rev:bool -> + match_rule -> + ((Id.t * context option) list * (* List of hypotheses matching: name + context *) + context option * (* Context for conclusion *) + Ltac_pretype.patvar_map (* Pattern variable substitution *)) Proofview.tactic diff --git a/vendor/Ltac2/src/tac2print.ml b/vendor/Ltac2/src/tac2print.ml new file mode 100644 index 0000000000..f4cb290265 --- /dev/null +++ b/vendor/Ltac2/src/tac2print.ml @@ -0,0 +1,488 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* str "'" ++ str (pr n) + | GTypRef (Other kn, []) -> pr_typref kn + | GTypRef (Other kn, [t]) -> + let paren = match lvl with + | T5_r | T5_l | T2 | T1 -> fun x -> x + | T0 -> paren + in + paren (pr_glbtype T1 t ++ spc () ++ pr_typref kn) + | GTypRef (Other kn, tl) -> + let paren = match lvl with + | T5_r | T5_l | T2 | T1 -> fun x -> x + | T0 -> paren + in + paren (str "(" ++ prlist_with_sep (fun () -> str ", ") (pr_glbtype lvl) tl ++ str ")" ++ spc () ++ pr_typref kn) + | GTypArrow (t1, t2) -> + let paren = match lvl with + | T5_r -> fun x -> x + | T5_l | T2 | T1 | T0 -> paren + in + paren (pr_glbtype T5_l t1 ++ spc () ++ str "->" ++ spc () ++ pr_glbtype T5_r t2) + | GTypRef (Tuple 0, []) -> + Libnames.pr_qualid (Tac2env.shortest_qualid_of_type t_unit) + | GTypRef (Tuple _, tl) -> + let paren = match lvl with + | T5_r | T5_l -> fun x -> x + | T2 | T1 | T0 -> paren + in + paren (prlist_with_sep (fun () -> str " * ") (pr_glbtype T2) tl) + in + hov 0 (pr_glbtype lvl c) + +let pr_glbtype pr c = pr_glbtype_gen pr T5_r c + +let int_name () = + let vars = ref Int.Map.empty in + fun n -> + if Int.Map.mem n !vars then Int.Map.find n !vars + else + let num = Int.Map.cardinal !vars in + let base = num mod 26 in + let rem = num / 26 in + let name = String.make 1 (Char.chr (97 + base)) in + let suff = if Int.equal rem 0 then "" else string_of_int rem in + let name = name ^ suff in + let () = vars := Int.Map.add n name !vars in + name + +(** Term printing *) + +let pr_constructor kn = + Libnames.pr_qualid (Tac2env.shortest_qualid_of_constructor kn) + +let pr_projection kn = + Libnames.pr_qualid (Tac2env.shortest_qualid_of_projection kn) + +type exp_level = Tac2expr.exp_level = +| E5 +| E4 +| E3 +| E2 +| E1 +| E0 + +let pr_atom = function +| AtmInt n -> Pp.int n +| AtmStr s -> qstring s + +let pr_name = function +| Name id -> Id.print id +| Anonymous -> str "_" + +let find_constructor n empty def = + let rec find n = function + | [] -> assert false + | (id, []) as ans :: rem -> + if empty then + if Int.equal n 0 then ans + else find (pred n) rem + else find n rem + | (id, _ :: _) as ans :: rem -> + if not empty then + if Int.equal n 0 then ans + else find (pred n) rem + else find n rem + in + find n def + +let pr_internal_constructor tpe n is_const = + let data = match Tac2env.interp_type tpe with + | (_, GTydAlg data) -> data + | _ -> assert false + in + let (id, _) = find_constructor n is_const data.galg_constructors in + let kn = change_kn_label tpe id in + pr_constructor kn + +let order_branches cbr nbr def = + let rec order cidx nidx def = match def with + | [] -> [] + | (id, []) :: rem -> + let ans = order (succ cidx) nidx rem in + (id, [], cbr.(cidx)) :: ans + | (id, _ :: _) :: rem -> + let ans = order cidx (succ nidx) rem in + let (vars, e) = nbr.(nidx) in + (id, Array.to_list vars, e) :: ans + in + order 0 0 def + +let pr_glbexpr_gen lvl c = + let rec pr_glbexpr lvl = function + | GTacAtm atm -> pr_atom atm + | GTacVar id -> Id.print id + | GTacRef gr -> + let qid = shortest_qualid_of_ltac (TacConstant gr) in + Libnames.pr_qualid qid + | GTacFun (nas, c) -> + let nas = pr_sequence pr_name nas in + let paren = match lvl with + | E0 | E1 | E2 | E3 | E4 -> paren + | E5 -> fun x -> x + in + paren (hov 0 (hov 2 (str "fun" ++ spc () ++ nas) ++ spc () ++ str "=>" ++ spc () ++ + pr_glbexpr E5 c)) + | GTacApp (c, cl) -> + let paren = match lvl with + | E0 -> paren + | E1 | E2 | E3 | E4 | E5 -> fun x -> x + in + paren (hov 2 (pr_glbexpr E1 c ++ spc () ++ (pr_sequence (pr_glbexpr E0) cl))) + | GTacLet (mut, bnd, e) -> + let paren = match lvl with + | E0 | E1 | E2 | E3 | E4 -> paren + | E5 -> fun x -> x + in + let mut = if mut then str "rec" ++ spc () else mt () in + let pr_bnd (na, e) = + pr_name na ++ spc () ++ str ":=" ++ spc () ++ hov 2 (pr_glbexpr E5 e) ++ spc () + in + let bnd = prlist_with_sep (fun () -> str "with" ++ spc ()) pr_bnd bnd in + paren (hv 0 (hov 2 (str "let" ++ spc () ++ mut ++ bnd ++ str "in") ++ spc () ++ pr_glbexpr E5 e)) + | GTacCst (Tuple 0, _, _) -> str "()" + | GTacCst (Tuple _, _, cl) -> + let paren = match lvl with + | E0 | E1 -> paren + | E2 | E3 | E4 | E5 -> fun x -> x + in + paren (prlist_with_sep (fun () -> str "," ++ spc ()) (pr_glbexpr E1) cl) + | GTacCst (Other tpe, n, cl) -> + pr_applied_constructor lvl tpe n cl + | GTacCse (e, info, cst_br, ncst_br) -> + let e = pr_glbexpr E5 e in + let br = match info with + | Other kn -> + let def = match Tac2env.interp_type kn with + | _, GTydAlg { galg_constructors = def } -> def + | _, GTydDef _ | _, GTydRec _ | _, GTydOpn -> assert false + in + let br = order_branches cst_br ncst_br def in + let pr_branch (cstr, vars, p) = + let cstr = change_kn_label kn cstr in + let cstr = pr_constructor cstr in + let vars = match vars with + | [] -> mt () + | _ -> spc () ++ pr_sequence pr_name vars + in + hov 4 (str "|" ++ spc () ++ hov 0 (cstr ++ vars ++ spc () ++ str "=>") ++ spc () ++ + hov 2 (pr_glbexpr E5 p)) ++ spc () + in + prlist pr_branch br + | Tuple n -> + let (vars, p) = if Int.equal n 0 then ([||], cst_br.(0)) else ncst_br.(0) in + let p = pr_glbexpr E5 p in + let vars = prvect_with_sep (fun () -> str "," ++ spc ()) pr_name vars in + hov 4 (str "|" ++ spc () ++ hov 0 (paren vars ++ spc () ++ str "=>") ++ spc () ++ p) + in + v 0 (hv 0 (str "match" ++ spc () ++ e ++ spc () ++ str "with") ++ spc () ++ br ++ spc () ++ str "end") + | GTacWth wth -> + let e = pr_glbexpr E5 wth.opn_match in + let pr_pattern c self vars p = + let self = match self with + | Anonymous -> mt () + | Name id -> spc () ++ str "as" ++ spc () ++ Id.print id + in + hov 4 (str "|" ++ spc () ++ hov 0 (c ++ vars ++ self ++ spc () ++ str "=>") ++ spc () ++ + hov 2 (pr_glbexpr E5 p)) ++ spc () + in + let pr_branch (cstr, (self, vars, p)) = + let cstr = pr_constructor cstr in + let vars = match Array.to_list vars with + | [] -> mt () + | vars -> spc () ++ pr_sequence pr_name vars + in + pr_pattern cstr self vars p + in + let br = prlist pr_branch (KNmap.bindings wth.opn_branch) in + let (def_as, def_p) = wth.opn_default in + let def = pr_pattern (str "_") def_as (mt ()) def_p in + let br = br ++ def in + v 0 (hv 0 (str "match" ++ spc () ++ e ++ spc () ++ str "with") ++ spc () ++ br ++ str "end") + | GTacPrj (kn, e, n) -> + let def = match Tac2env.interp_type kn with + | _, GTydRec def -> def + | _, GTydDef _ | _, GTydAlg _ | _, GTydOpn -> assert false + in + let (proj, _, _) = List.nth def n in + let proj = change_kn_label kn proj in + let proj = pr_projection proj in + let e = pr_glbexpr E0 e in + hov 0 (e ++ str "." ++ paren proj) + | GTacSet (kn, e, n, r) -> + let def = match Tac2env.interp_type kn with + | _, GTydRec def -> def + | _, GTydDef _ | _, GTydAlg _ | _, GTydOpn -> assert false + in + let (proj, _, _) = List.nth def n in + let proj = change_kn_label kn proj in + let proj = pr_projection proj in + let e = pr_glbexpr E0 e in + let r = pr_glbexpr E1 r in + hov 0 (e ++ str "." ++ paren proj ++ spc () ++ str ":=" ++ spc () ++ r) + | GTacOpn (kn, cl) -> + let paren = match lvl with + | E0 -> paren + | E1 | E2 | E3 | E4 | E5 -> fun x -> x + in + let c = pr_constructor kn in + paren (hov 0 (c ++ spc () ++ (pr_sequence (pr_glbexpr E0) cl))) + | GTacExt (tag, arg) -> + let tpe = interp_ml_object tag in + hov 0 (tpe.ml_print (Global.env ()) arg) (* FIXME *) + | GTacPrm (prm, args) -> + let args = match args with + | [] -> mt () + | _ -> spc () ++ pr_sequence (pr_glbexpr E0) args + in + hov 0 (str "@external" ++ spc () ++ qstring prm.mltac_plugin ++ spc () ++ + qstring prm.mltac_tactic ++ args) + and pr_applied_constructor lvl tpe n cl = + let _, data = Tac2env.interp_type tpe in + if KerName.equal tpe t_list then + let rec factorize accu = function + | GTacCst (_, 0, []) -> accu, None + | GTacCst (_, 0, [e; l]) -> factorize (e :: accu) l + | e -> accu, Some e + in + let l, e = factorize [] (GTacCst (Other tpe, n, cl)) in + match e with + | None -> + let pr e = pr_glbexpr E4 e in + hov 2 (str "[" ++ prlist_with_sep pr_semicolon pr (List.rev l) ++ str "]") + | Some e -> + let paren = match lvl with + | E0 | E1 | E2 -> paren + | E3 | E4 | E5 -> fun x -> x + in + let pr e = pr_glbexpr E1 e in + let pr_cons () = spc () ++ str "::" ++ spc () in + paren (hov 2 (prlist_with_sep pr_cons pr (List.rev (e :: l)))) + else match data with + | GTydAlg def -> + let paren = match lvl with + | E0 -> + if List.is_empty cl then fun x -> x else paren + | E1 | E2 | E3 | E4 | E5 -> fun x -> x + in + let cstr = pr_internal_constructor tpe n (List.is_empty cl) in + let cl = match cl with + | [] -> mt () + | _ -> spc () ++ pr_sequence (pr_glbexpr E0) cl + in + paren (hov 2 (cstr ++ cl)) + | GTydRec def -> + let args = List.combine def cl in + let pr_arg ((id, _, _), arg) = + let kn = change_kn_label tpe id in + pr_projection kn ++ spc () ++ str ":=" ++ spc () ++ pr_glbexpr E1 arg + in + let args = prlist_with_sep pr_semicolon pr_arg args in + hv 0 (str "{" ++ spc () ++ args ++ spc () ++ str "}") + | (GTydDef _ | GTydOpn) -> assert false + in + hov 0 (pr_glbexpr lvl c) + + + +let pr_glbexpr c = + pr_glbexpr_gen E5 c + +(** Toplevel printers *) + +let rec subst_type subst (t : 'a glb_typexpr) = match t with +| GTypVar id -> subst.(id) +| GTypArrow (t1, t2) -> GTypArrow (subst_type subst t1, subst_type subst t2) +| GTypRef (qid, args) -> + GTypRef (qid, List.map (fun t -> subst_type subst t) args) + +let unfold kn args = + let (nparams, def) = Tac2env.interp_type kn in + match def with + | GTydDef (Some def) -> + let args = Array.of_list args in + Some (subst_type args def) + | _ -> None + +let rec kind t = match t with +| GTypVar id -> GTypVar id +| GTypRef (Other kn, tl) -> + begin match unfold kn tl with + | None -> t + | Some t -> kind t + end +| GTypArrow _ | GTypRef (Tuple _, _) -> t + +type val_printer = + { val_printer : 'a. Environ.env -> Evd.evar_map -> valexpr -> 'a glb_typexpr list -> Pp.t } + +let printers = ref KNmap.empty + +let register_val_printer kn pr = + printers := KNmap.add kn pr !printers + +open Tac2ffi + +let rec pr_valexpr env sigma v t = match kind t with +| GTypVar _ -> str "" +| GTypRef (Other kn, params) -> + let pr = try Some (KNmap.find kn !printers) with Not_found -> None in + begin match pr with + | Some pr -> pr.val_printer env sigma v params + | None -> + let n, repr = Tac2env.interp_type kn in + if KerName.equal kn t_list then + pr_val_list env sigma (to_list (fun v -> repr_to valexpr v) v) (List.hd params) + else match repr with + | GTydDef None -> str "" + | GTydDef (Some _) -> + (* Shouldn't happen thanks to kind *) + assert false + | GTydAlg alg -> + if Valexpr.is_int v then + pr_internal_constructor kn (Tac2ffi.to_int v) true + else + let (n, args) = Tac2ffi.to_block v in + let (id, tpe) = find_constructor n false alg.galg_constructors in + let knc = change_kn_label kn id in + let args = pr_constrargs env sigma params args tpe in + hv 2 (pr_constructor knc ++ spc () ++ str "(" ++ args ++ str ")") + | GTydRec rcd -> + let (_, args) = Tac2ffi.to_block v in + pr_record env sigma params args rcd + | GTydOpn -> + begin match Tac2ffi.to_open v with + | (knc, [||]) -> pr_constructor knc + | (knc, args) -> + let data = Tac2env.interp_constructor knc in + let args = pr_constrargs env sigma params args data.Tac2env.cdata_args in + hv 2 (pr_constructor knc ++ spc () ++ str "(" ++ args ++ str ")") + end + end +| GTypArrow _ -> str "" +| GTypRef (Tuple 0, []) -> str "()" +| GTypRef (Tuple _, tl) -> + let blk = Array.to_list (snd (to_block v)) in + if List.length blk == List.length tl then + let prs = List.map2 (fun v t -> pr_valexpr env sigma v t) blk tl in + hv 2 (str "(" ++ prlist_with_sep pr_comma (fun p -> p) prs ++ str ")") + else + str "" + +and pr_constrargs env sigma params args tpe = + let subst = Array.of_list params in + let tpe = List.map (fun t -> subst_type subst t) tpe in + let args = Array.to_list args in + let args = List.combine args tpe in + prlist_with_sep pr_comma (fun (v, t) -> pr_valexpr env sigma v t) args + +and pr_record env sigma params args rcd = + let subst = Array.of_list params in + let map (id, _, tpe) = (id, subst_type subst tpe) in + let rcd = List.map map rcd in + let args = Array.to_list args in + let fields = List.combine rcd args in + let pr_field ((id, t), arg) = + Id.print id ++ spc () ++ str ":=" ++ spc () ++ pr_valexpr env sigma arg t + in + str "{" ++ spc () ++ prlist_with_sep pr_semicolon pr_field fields ++ spc () ++ str "}" + +and pr_val_list env sigma args tpe = + let pr v = pr_valexpr env sigma v tpe in + str "[" ++ prlist_with_sep pr_semicolon pr args ++ str "]" + +let register_init n f = + let kn = KerName.make Tac2env.coq_prefix (Label.make n) in + register_val_printer kn { val_printer = fun env sigma v _ -> f env sigma v } + +let () = register_init "int" begin fun _ _ n -> + let n = to_int n in + Pp.int n +end + +let () = register_init "string" begin fun _ _ s -> + let s = to_string s in + Pp.quote (str (Bytes.to_string s)) +end + +let () = register_init "ident" begin fun _ _ id -> + let id = to_ident id in + str "@" ++ Id.print id +end + +let () = register_init "constr" begin fun env sigma c -> + let c = to_constr c in + let c = try Printer.pr_leconstr_env env sigma c with _ -> str "..." in + str "constr:(" ++ c ++ str ")" +end + +let () = register_init "pattern" begin fun env sigma c -> + let c = to_pattern c in + let c = try Printer.pr_lconstr_pattern_env env sigma c with _ -> str "..." in + str "pattern:(" ++ c ++ str ")" +end + +let () = register_init "message" begin fun _ _ pp -> + str "message:(" ++ to_pp pp ++ str ")" +end + +let () = register_init "err" begin fun _ _ e -> + let e = to_ext val_exn e in + let (e, _) = ExplainErr.process_vernac_interp_error ~allow_uncaught:true e in + str "err:(" ++ CErrors.print_no_report e ++ str ")" +end + +let () = + let kn = KerName.make Tac2env.coq_prefix (Label.make "array") in + let val_printer env sigma v arg = match arg with + | [arg] -> + let (_, v) = to_block v in + str "[|" ++ spc () ++ + prvect_with_sep pr_semicolon (fun a -> pr_valexpr env sigma a arg) v ++ + spc () ++ str "|]" + | _ -> assert false + in + register_val_printer kn { val_printer } diff --git a/vendor/Ltac2/src/tac2print.mli b/vendor/Ltac2/src/tac2print.mli new file mode 100644 index 0000000000..9b9db2937d --- /dev/null +++ b/vendor/Ltac2/src/tac2print.mli @@ -0,0 +1,46 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Pp.t +val pr_glbtype_gen : ('a -> string) -> typ_level -> 'a glb_typexpr -> Pp.t +val pr_glbtype : ('a -> string) -> 'a glb_typexpr -> Pp.t + +(** {5 Printing expressions} *) + +val pr_constructor : ltac_constructor -> Pp.t +val pr_internal_constructor : type_constant -> int -> bool -> Pp.t +val pr_projection : ltac_projection -> Pp.t +val pr_glbexpr_gen : exp_level -> glb_tacexpr -> Pp.t +val pr_glbexpr : glb_tacexpr -> Pp.t + +(** {5 Printing values}*) + +type val_printer = + { val_printer : 'a. Environ.env -> Evd.evar_map -> valexpr -> 'a glb_typexpr list -> Pp.t } + +val register_val_printer : type_constant -> val_printer -> unit + +val pr_valexpr : Environ.env -> Evd.evar_map -> valexpr -> 'a glb_typexpr -> Pp.t + +(** {5 Utilities} *) + +val int_name : unit -> (int -> string) +(** Create a function that give names to integers. The names are generated on + the fly, in the order they are encountered. *) diff --git a/vendor/Ltac2/src/tac2qexpr.mli b/vendor/Ltac2/src/tac2qexpr.mli new file mode 100644 index 0000000000..400ab1a092 --- /dev/null +++ b/vendor/Ltac2/src/tac2qexpr.mli @@ -0,0 +1,173 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* + CAst.make ?loc @@ CTacCst (AbsKn (Tuple 0)) +| [e] -> e +| el -> + let len = List.length el in + CAst.make ?loc @@ CTacApp (CAst.make ?loc @@ CTacCst (AbsKn (Tuple len)), el) + +let of_int {loc;v=n} = + CAst.make ?loc @@ CTacAtm (AtmInt n) + +let of_option ?loc f opt = match opt with +| None -> constructor ?loc (coq_core "None") [] +| Some e -> constructor ?loc (coq_core "Some") [f e] + +let inj_wit ?loc wit x = + CAst.make ?loc @@ CTacExt (wit, x) + +let of_variable {loc;v=id} = + let qid = Libnames.qualid_of_ident ?loc id in + if Tac2env.is_constructor qid then + CErrors.user_err ?loc (str "Invalid identifier") + else CAst.make ?loc @@ CTacRef (RelId qid) + +let of_anti f = function +| QExpr x -> f x +| QAnti id -> of_variable id + +let of_ident {loc;v=id} = inj_wit ?loc wit_ident id + +let of_constr c = + let loc = Constrexpr_ops.constr_loc c in + inj_wit ?loc wit_constr c + +let of_open_constr c = + let loc = Constrexpr_ops.constr_loc c in + inj_wit ?loc wit_open_constr c + +let of_bool ?loc b = + let c = if b then coq_core "true" else coq_core "false" in + constructor ?loc c [] + +let rec of_list ?loc f = function +| [] -> constructor (coq_core "[]") [] +| e :: l -> + constructor ?loc (coq_core "::") [f e; of_list ?loc f l] + +let of_qhyp {loc;v=h} = match h with +| QAnonHyp n -> std_constructor ?loc "AnonHyp" [of_int n] +| QNamedHyp id -> std_constructor ?loc "NamedHyp" [of_ident id] + +let of_bindings {loc;v=b} = match b with +| QNoBindings -> + std_constructor ?loc "NoBindings" [] +| QImplicitBindings tl -> + std_constructor ?loc "ImplicitBindings" [of_list ?loc of_open_constr tl] +| QExplicitBindings tl -> + let map e = of_pair (fun q -> of_anti of_qhyp q) of_open_constr e in + std_constructor ?loc "ExplicitBindings" [of_list ?loc map tl] + +let of_constr_with_bindings c = of_pair of_open_constr of_bindings c + +let rec of_intro_pattern {loc;v=pat} = match pat with +| QIntroForthcoming b -> + std_constructor ?loc "IntroForthcoming" [of_bool b] +| QIntroNaming iname -> + std_constructor ?loc "IntroNaming" [of_intro_pattern_naming iname] +| QIntroAction iact -> + std_constructor ?loc "IntroAction" [of_intro_pattern_action iact] + +and of_intro_pattern_naming {loc;v=pat} = match pat with +| QIntroIdentifier id -> + std_constructor ?loc "IntroIdentifier" [of_anti of_ident id] +| QIntroFresh id -> + std_constructor ?loc "IntroFresh" [of_anti of_ident id] +| QIntroAnonymous -> + std_constructor ?loc "IntroAnonymous" [] + +and of_intro_pattern_action {loc;v=pat} = match pat with +| QIntroWildcard -> + std_constructor ?loc "IntroWildcard" [] +| QIntroOrAndPattern pat -> + std_constructor ?loc "IntroOrAndPattern" [of_or_and_intro_pattern pat] +| QIntroInjection il -> + std_constructor ?loc "IntroInjection" [of_intro_patterns il] +| QIntroRewrite b -> + std_constructor ?loc "IntroRewrite" [of_bool ?loc b] + +and of_or_and_intro_pattern {loc;v=pat} = match pat with +| QIntroOrPattern ill -> + std_constructor ?loc "IntroOrPattern" [of_list ?loc of_intro_patterns ill] +| QIntroAndPattern il -> + std_constructor ?loc "IntroAndPattern" [of_intro_patterns il] + +and of_intro_patterns {loc;v=l} = + of_list ?loc of_intro_pattern l + +let of_hyp_location_flag ?loc = function +| Locus.InHyp -> std_constructor ?loc "InHyp" [] +| Locus.InHypTypeOnly -> std_constructor ?loc "InHypTypeOnly" [] +| Locus.InHypValueOnly -> std_constructor ?loc "InHypValueOnly" [] + +let of_occurrences {loc;v=occ} = match occ with +| QAllOccurrences -> std_constructor ?loc "AllOccurrences" [] +| QAllOccurrencesBut occs -> + let map occ = of_anti of_int occ in + let occs = of_list ?loc map occs in + std_constructor ?loc "AllOccurrencesBut" [occs] +| QNoOccurrences -> std_constructor ?loc "NoOccurrences" [] +| QOnlyOccurrences occs -> + let map occ = of_anti of_int occ in + let occs = of_list ?loc map occs in + std_constructor ?loc "OnlyOccurrences" [occs] + +let of_hyp_location ?loc ((occs, id), flag) = + of_tuple ?loc [ + of_anti of_ident id; + of_occurrences occs; + of_hyp_location_flag ?loc flag; + ] + +let of_clause {loc;v=cl} = + let hyps = of_option ?loc (fun l -> of_list ?loc of_hyp_location l) cl.q_onhyps in + let concl = of_occurrences cl.q_concl_occs in + CAst.make ?loc @@ CTacRec ([ + std_proj "on_hyps", hyps; + std_proj "on_concl", concl; + ]) + +let of_destruction_arg {loc;v=arg} = match arg with +| QElimOnConstr c -> + let arg = thunk (of_constr_with_bindings c) in + std_constructor ?loc "ElimOnConstr" [arg] +| QElimOnIdent id -> std_constructor ?loc "ElimOnIdent" [of_ident id] +| QElimOnAnonHyp n -> std_constructor ?loc "ElimOnAnonHyp" [of_int n] + +let of_induction_clause {loc;v=cl} = + let arg = of_destruction_arg cl.indcl_arg in + let eqn = of_option ?loc of_intro_pattern_naming cl.indcl_eqn in + let as_ = of_option ?loc of_or_and_intro_pattern cl.indcl_as in + let in_ = of_option ?loc of_clause cl.indcl_in in + CAst.make ?loc @@ CTacRec ([ + std_proj "indcl_arg", arg; + std_proj "indcl_eqn", eqn; + std_proj "indcl_as", as_; + std_proj "indcl_in", in_; + ]) + +let check_pattern_id ?loc id = + if Tac2env.is_constructor (Libnames.qualid_of_ident id) then + CErrors.user_err ?loc (str "Invalid pattern binding name " ++ Id.print id) + +let pattern_vars pat = + let rec aux () accu pat = match pat.CAst.v with + | Constrexpr.CPatVar id + | Constrexpr.CEvar (id, []) -> + let () = check_pattern_id ?loc:pat.CAst.loc id in + Id.Set.add id accu + | _ -> + Constrexpr_ops.fold_constr_expr_with_binders (fun _ () -> ()) aux () accu pat + in + aux () Id.Set.empty pat + +let abstract_vars loc vars tac = + let get_name = function Name id -> Some id | Anonymous -> None in + let def = try Some (List.find_map get_name vars) with Not_found -> None in + let na, tac = match def with + | None -> (Anonymous, tac) + | Some id0 -> + (* Trick: in order not to shadow a variable nor to choose an arbitrary + name, we reuse one which is going to be shadowed by the matched + variables anyways. *) + let build_bindings (n, accu) na = match na with + | Anonymous -> (n + 1, accu) + | Name _ -> + let get = global_ref ?loc (kername array_prefix "get") in + let args = [of_variable CAst.(make ?loc id0); of_int CAst.(make ?loc n)] in + let e = CAst.make ?loc @@ CTacApp (get, args) in + let accu = (CAst.make ?loc @@ CPatVar na, e) :: accu in + (n + 1, accu) + in + let (_, bnd) = List.fold_left build_bindings (0, []) vars in + let tac = CAst.make ?loc @@ CTacLet (false, bnd, tac) in + (Name id0, tac) + in + CAst.make ?loc @@ CTacFun ([CAst.make ?loc @@ CPatVar na], tac) + +let of_pattern p = + inj_wit ?loc:p.CAst.loc wit_pattern p + +let of_conversion {loc;v=c} = match c with +| QConvert c -> + let pat = of_option ?loc of_pattern None in + let c = CAst.make ?loc @@ CTacFun ([CAst.make ?loc @@ CPatVar Anonymous], of_constr c) in + of_tuple ?loc [pat; c] +| QConvertWith (pat, c) -> + let vars = pattern_vars pat in + let pat = of_option ?loc of_pattern (Some pat) in + let c = of_constr c in + (* Order is critical here *) + let vars = List.map (fun id -> Name id) (Id.Set.elements vars) in + let c = abstract_vars loc vars c in + of_tuple [pat; c] + +let of_repeat {loc;v=r} = match r with +| QPrecisely n -> std_constructor ?loc "Precisely" [of_int n] +| QUpTo n -> std_constructor ?loc "UpTo" [of_int n] +| QRepeatStar -> std_constructor ?loc "RepeatStar" [] +| QRepeatPlus -> std_constructor ?loc "RepeatPlus" [] + +let of_orient loc b = + if b then std_constructor ?loc "LTR" [] + else std_constructor ?loc "RTL" [] + +let of_rewriting {loc;v=rew} = + let orient = + let {loc;v=orient} = rew.rew_orient in + of_option ?loc (fun b -> of_orient loc b) orient + in + let repeat = of_repeat rew.rew_repeat in + let equatn = thunk (of_constr_with_bindings rew.rew_equatn) in + CAst.make ?loc @@ CTacRec ([ + std_proj "rew_orient", orient; + std_proj "rew_repeat", repeat; + std_proj "rew_equatn", equatn; + ]) + +let of_hyp ?loc id = + let hyp = global_ref ?loc (control_core "hyp") in + CAst.make ?loc @@ CTacApp (hyp, [of_ident id]) + +let of_exact_hyp ?loc id = + let refine = global_ref ?loc (control_core "refine") in + CAst.make ?loc @@ CTacApp (refine, [thunk (of_hyp ?loc id)]) + +let of_exact_var ?loc id = + let refine = global_ref ?loc (control_core "refine") in + CAst.make ?loc @@ CTacApp (refine, [thunk (of_variable id)]) + +let of_dispatch tacs = + let loc = tacs.loc in + let default = function + | Some e -> thunk e + | None -> thunk (CAst.make ?loc @@ CTacCst (AbsKn (Tuple 0))) + in + let map e = of_pair default (fun l -> of_list ?loc default l) (CAst.make ?loc e) in + of_pair (fun l -> of_list ?loc default l) (fun r -> of_option ?loc map r) tacs + +let make_red_flag l = + let open Genredexpr in + let rec add_flag red = function + | [] -> red + | {v=flag} :: lf -> + let red = match flag with + | QBeta -> { red with rBeta = true } + | QMatch -> { red with rMatch = true } + | QFix -> { red with rFix = true } + | QCofix -> { red with rCofix = true } + | QZeta -> { red with rZeta = true } + | QConst {loc;v=l} -> + if red.rDelta then + CErrors.user_err ?loc Pp.(str + "Cannot set both constants to unfold and constants not to unfold"); + { red with rConst = red.rConst @ l } + | QDeltaBut {loc;v=l} -> + if red.rConst <> [] && not red.rDelta then + CErrors.user_err ?loc Pp.(str + "Cannot set both constants to unfold and constants not to unfold"); + { red with rConst = red.rConst @ l; rDelta = true } + | QIota -> + { red with rMatch = true; rFix = true; rCofix = true } + in + add_flag red lf + in + add_flag + {rBeta = false; rMatch = false; rFix = false; rCofix = false; + rZeta = false; rDelta = false; rConst = []} + l + +let of_reference r = + let of_ref ref = + inj_wit ?loc:ref.loc wit_reference ref + in + of_anti of_ref r + +let of_strategy_flag {loc;v=flag} = + let open Genredexpr in + let flag = make_red_flag flag in + CAst.make ?loc @@ CTacRec ([ + std_proj "rBeta", of_bool ?loc flag.rBeta; + std_proj "rMatch", of_bool ?loc flag.rMatch; + std_proj "rFix", of_bool ?loc flag.rFix; + std_proj "rCofix", of_bool ?loc flag.rCofix; + std_proj "rZeta", of_bool ?loc flag.rZeta; + std_proj "rDelta", of_bool ?loc flag.rDelta; + std_proj "rConst", of_list ?loc of_reference flag.rConst; + ]) + +let of_hintdb {loc;v=hdb} = match hdb with +| QHintAll -> of_option ?loc (fun l -> of_list (fun id -> of_anti of_ident id) l) None +| QHintDbs ids -> of_option ?loc (fun l -> of_list (fun id -> of_anti of_ident id) l) (Some ids) + +let extract_name ?loc oid = match oid with +| None -> Anonymous +| Some id -> + let () = check_pattern_id ?loc id in + Name id + +(** For every branch in the matching, generate a corresponding term of type + [(match_kind * pattern * (context -> constr array -> 'a))] + where the function binds the names from the pattern to the contents of the + constr array. *) +let of_constr_matching {loc;v=m} = + let map {loc;v=({loc=ploc;v=pat}, tac)} = + let (knd, pat, na) = match pat with + | QConstrMatchPattern pat -> + let knd = constructor ?loc (pattern_core "MatchPattern") [] in + (knd, pat, Anonymous) + | QConstrMatchContext (id, pat) -> + let na = extract_name ?loc id in + let knd = constructor ?loc (pattern_core "MatchContext") [] in + (knd, pat, na) + in + let vars = pattern_vars pat in + (* Order of elements is crucial here! *) + let vars = Id.Set.elements vars in + let vars = List.map (fun id -> Name id) vars in + let e = abstract_vars loc vars tac in + let e = CAst.make ?loc @@ CTacFun ([CAst.make ?loc @@ CPatVar na], e) in + let pat = inj_wit ?loc:ploc wit_pattern pat in + of_tuple [knd; pat; e] + in + of_list ?loc map m + +(** From the patterns and the body of the branch, generate: + - a goal pattern: (constr_match list * constr_match) + - a branch function (ident array -> context array -> constr array -> context -> 'a) +*) +let of_goal_matching {loc;v=gm} = + let mk_pat {loc;v=p} = match p with + | QConstrMatchPattern pat -> + let knd = constructor ?loc (pattern_core "MatchPattern") [] in + (Anonymous, pat, knd) + | QConstrMatchContext (id, pat) -> + let na = extract_name ?loc id in + let knd = constructor ?loc (pattern_core "MatchContext") [] in + (na, pat, knd) + in + let mk_gpat {loc;v=p} = + let concl_pat = p.q_goal_match_concl in + let hyps_pats = p.q_goal_match_hyps in + let (concl_ctx, concl_pat, concl_knd) = mk_pat concl_pat in + let vars = pattern_vars concl_pat in + let map accu (na, pat) = + let (ctx, pat, knd) = mk_pat pat in + let vars = pattern_vars pat in + (Id.Set.union vars accu, (na, ctx, pat, knd)) + in + let (vars, hyps_pats) = List.fold_left_map map vars hyps_pats in + let map (_, _, pat, knd) = of_tuple [knd; of_pattern pat] in + let concl = of_tuple [concl_knd; of_pattern concl_pat] in + let r = of_tuple [of_list ?loc map hyps_pats; concl] in + let hyps = List.map (fun ({CAst.v=na}, _, _, _) -> na) hyps_pats in + let map (_, na, _, _) = na in + let hctx = List.map map hyps_pats in + (* Order of elements is crucial here! *) + let vars = Id.Set.elements vars in + let subst = List.map (fun id -> Name id) vars in + (r, hyps, hctx, subst, concl_ctx) + in + let map {loc;v=(pat, tac)} = + let (pat, hyps, hctx, subst, cctx) = mk_gpat pat in + let tac = CAst.make ?loc @@ CTacFun ([CAst.make ?loc @@ CPatVar cctx], tac) in + let tac = abstract_vars loc subst tac in + let tac = abstract_vars loc hctx tac in + let tac = abstract_vars loc hyps tac in + of_tuple ?loc [pat; tac] + in + of_list ?loc map gm + +let of_move_location {loc;v=mv} = match mv with +| QMoveAfter id -> std_constructor ?loc "MoveAfter" [of_anti of_ident id] +| QMoveBefore id -> std_constructor ?loc "MoveBefore" [of_anti of_ident id] +| QMoveFirst -> std_constructor ?loc "MoveFirst" [] +| QMoveLast -> std_constructor ?loc "MoveLast" [] + +let of_pose p = + of_pair (fun id -> of_option (fun id -> of_anti of_ident id) id) of_open_constr p + +let of_assertion {loc;v=ast} = match ast with +| QAssertType (ipat, c, tac) -> + let ipat = of_option of_intro_pattern ipat in + let c = of_constr c in + let tac = of_option thunk tac in + std_constructor ?loc "AssertType" [ipat; c; tac] +| QAssertValue (id, c) -> + let id = of_anti of_ident id in + let c = of_constr c in + std_constructor ?loc "AssertValue" [id; c] diff --git a/vendor/Ltac2/src/tac2quote.mli b/vendor/Ltac2/src/tac2quote.mli new file mode 100644 index 0000000000..1b03dad8ec --- /dev/null +++ b/vendor/Ltac2/src/tac2quote.mli @@ -0,0 +1,102 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* ltac_constructor -> raw_tacexpr list -> raw_tacexpr + +val thunk : raw_tacexpr -> raw_tacexpr + +val of_anti : ('a -> raw_tacexpr) -> 'a or_anti -> raw_tacexpr + +val of_int : int CAst.t -> raw_tacexpr + +val of_pair : ('a -> raw_tacexpr) -> ('b -> raw_tacexpr) -> ('a * 'b) CAst.t -> raw_tacexpr + +val of_tuple : ?loc:Loc.t -> raw_tacexpr list -> raw_tacexpr + +val of_variable : Id.t CAst.t -> raw_tacexpr + +val of_ident : Id.t CAst.t -> raw_tacexpr + +val of_constr : Constrexpr.constr_expr -> raw_tacexpr + +val of_open_constr : Constrexpr.constr_expr -> raw_tacexpr + +val of_list : ?loc:Loc.t -> ('a -> raw_tacexpr) -> 'a list -> raw_tacexpr + +val of_bindings : bindings -> raw_tacexpr + +val of_intro_pattern : intro_pattern -> raw_tacexpr + +val of_intro_patterns : intro_pattern list CAst.t -> raw_tacexpr + +val of_clause : clause -> raw_tacexpr + +val of_destruction_arg : destruction_arg -> raw_tacexpr + +val of_induction_clause : induction_clause -> raw_tacexpr + +val of_conversion : conversion -> raw_tacexpr + +val of_rewriting : rewriting -> raw_tacexpr + +val of_occurrences : occurrences -> raw_tacexpr + +val of_hintdb : hintdb -> raw_tacexpr + +val of_move_location : move_location -> raw_tacexpr + +val of_reference : reference or_anti -> raw_tacexpr + +val of_hyp : ?loc:Loc.t -> Id.t CAst.t -> raw_tacexpr +(** id ↦ 'Control.hyp @id' *) + +val of_exact_hyp : ?loc:Loc.t -> Id.t CAst.t -> raw_tacexpr +(** id ↦ 'Control.refine (fun () => Control.hyp @id') *) + +val of_exact_var : ?loc:Loc.t -> Id.t CAst.t -> raw_tacexpr +(** id ↦ 'Control.refine (fun () => Control.hyp @id') *) + +val of_dispatch : dispatch -> raw_tacexpr + +val of_strategy_flag : strategy_flag -> raw_tacexpr + +val of_pose : pose -> raw_tacexpr + +val of_assertion : assertion -> raw_tacexpr + +val of_constr_matching : constr_matching -> raw_tacexpr + +val of_goal_matching : goal_matching -> raw_tacexpr + +(** {5 Generic arguments} *) + +val wit_pattern : (Constrexpr.constr_expr, Pattern.constr_pattern) Arg.tag + +val wit_ident : (Id.t, Id.t) Arg.tag + +val wit_reference : (reference, GlobRef.t) Arg.tag + +val wit_constr : (Constrexpr.constr_expr, Glob_term.glob_constr) Arg.tag + +val wit_open_constr : (Constrexpr.constr_expr, Glob_term.glob_constr) Arg.tag + +val wit_ltac1 : (Ltac_plugin.Tacexpr.raw_tactic_expr, Ltac_plugin.Tacexpr.glob_tactic_expr) Arg.tag +(** Ltac1 AST quotation, seen as a 'tactic'. Its type is unit in Ltac2. *) + +val wit_ltac1val : (Ltac_plugin.Tacexpr.raw_tactic_expr, Ltac_plugin.Tacexpr.glob_tactic_expr) Arg.tag +(** Ltac1 AST quotation, seen as a value-returning expression, with type Ltac1.t. *) diff --git a/vendor/Ltac2/src/tac2stdlib.ml b/vendor/Ltac2/src/tac2stdlib.ml new file mode 100644 index 0000000000..ffef2c05fd --- /dev/null +++ b/vendor/Ltac2/src/tac2stdlib.ml @@ -0,0 +1,578 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* assert false) f + +let return x = Proofview.tclUNIT x +let v_unit = Value.of_unit () +let thaw r f = Tac2ffi.app_fun1 f unit r () +let uthaw r f = Tac2ffi.app_fun1 (to_fun1 unit r f) unit r () +let thunk r = fun1 unit r + +let to_name c = match Value.to_option Value.to_ident c with +| None -> Anonymous +| Some id -> Name id + +let name = make_to_repr to_name + +let to_occurrences = function +| ValInt 0 -> AllOccurrences +| ValBlk (0, [| vl |]) -> AllOccurrencesBut (Value.to_list Value.to_int vl) +| ValInt 1 -> NoOccurrences +| ValBlk (1, [| vl |]) -> OnlyOccurrences (Value.to_list Value.to_int vl) +| _ -> assert false + +let occurrences = make_to_repr to_occurrences + +let to_hyp_location_flag v = match Value.to_int v with +| 0 -> InHyp +| 1 -> InHypTypeOnly +| 2 -> InHypValueOnly +| _ -> assert false + +let to_clause v = match Value.to_tuple v with +| [| hyps; concl |] -> + let cast v = match Value.to_tuple v with + | [| hyp; occ; flag |] -> + (Value.to_ident hyp, to_occurrences occ, to_hyp_location_flag flag) + | _ -> assert false + in + let hyps = Value.to_option (fun h -> Value.to_list cast h) hyps in + { onhyps = hyps; concl_occs = to_occurrences concl; } +| _ -> assert false + +let clause = make_to_repr to_clause + +let to_red_flag v = match Value.to_tuple v with +| [| beta; iota; fix; cofix; zeta; delta; const |] -> + { + rBeta = Value.to_bool beta; + rMatch = Value.to_bool iota; + rFix = Value.to_bool fix; + rCofix = Value.to_bool cofix; + rZeta = Value.to_bool zeta; + rDelta = Value.to_bool delta; + rConst = Value.to_list Value.to_reference const; + } +| _ -> assert false + +let red_flags = make_to_repr to_red_flag + +let pattern_with_occs = pair pattern occurrences + +let constr_with_occs = pair constr occurrences + +let reference_with_occs = pair reference occurrences + +let rec to_intro_pattern v = match Value.to_block v with +| (0, [| b |]) -> IntroForthcoming (Value.to_bool b) +| (1, [| pat |]) -> IntroNaming (to_intro_pattern_naming pat) +| (2, [| act |]) -> IntroAction (to_intro_pattern_action act) +| _ -> assert false + +and to_intro_pattern_naming = function +| ValBlk (0, [| id |]) -> IntroIdentifier (Value.to_ident id) +| ValBlk (1, [| id |]) -> IntroFresh (Value.to_ident id) +| ValInt 0 -> IntroAnonymous +| _ -> assert false + +and to_intro_pattern_action = function +| ValInt 0 -> IntroWildcard +| ValBlk (0, [| op |]) -> IntroOrAndPattern (to_or_and_intro_pattern op) +| ValBlk (1, [| inj |]) -> + let map ipat = to_intro_pattern ipat in + IntroInjection (Value.to_list map inj) +| ValBlk (2, [| c; ipat |]) -> + let c = Value.to_fun1 Value.unit Value.constr c in + IntroApplyOn (c, to_intro_pattern ipat) +| ValBlk (3, [| b |]) -> IntroRewrite (Value.to_bool b) +| _ -> assert false + +and to_or_and_intro_pattern v = match Value.to_block v with +| (0, [| ill |]) -> + IntroOrPattern (Value.to_list to_intro_patterns ill) +| (1, [| il |]) -> + IntroAndPattern (to_intro_patterns il) +| _ -> assert false + +and to_intro_patterns il = + Value.to_list to_intro_pattern il + +let intro_pattern = make_to_repr to_intro_pattern + +let intro_patterns = make_to_repr to_intro_patterns + +let to_destruction_arg v = match Value.to_block v with +| (0, [| c |]) -> + let c = uthaw constr_with_bindings c in + ElimOnConstr c +| (1, [| id |]) -> ElimOnIdent (Value.to_ident id) +| (2, [| n |]) -> ElimOnAnonHyp (Value.to_int n) +| _ -> assert false + +let destruction_arg = make_to_repr to_destruction_arg + +let to_induction_clause v = match Value.to_tuple v with +| [| arg; eqn; as_; in_ |] -> + let arg = to_destruction_arg arg in + let eqn = Value.to_option to_intro_pattern_naming eqn in + let as_ = Value.to_option to_or_and_intro_pattern as_ in + let in_ = Value.to_option to_clause in_ in + (arg, eqn, as_, in_) +| _ -> + assert false + +let induction_clause = make_to_repr to_induction_clause + +let to_assertion v = match Value.to_block v with +| (0, [| ipat; t; tac |]) -> + let to_tac t = Value.to_fun1 Value.unit Value.unit t in + let ipat = Value.to_option to_intro_pattern ipat in + let t = Value.to_constr t in + let tac = Value.to_option to_tac tac in + AssertType (ipat, t, tac) +| (1, [| id; c |]) -> + AssertValue (Value.to_ident id, Value.to_constr c) +| _ -> assert false + +let assertion = make_to_repr to_assertion + +let to_multi = function +| ValBlk (0, [| n |]) -> Precisely (Value.to_int n) +| ValBlk (1, [| n |]) -> UpTo (Value.to_int n) +| ValInt 0 -> RepeatStar +| ValInt 1 -> RepeatPlus +| _ -> assert false + +let to_rewriting v = match Value.to_tuple v with +| [| orient; repeat; c |] -> + let orient = Value.to_option Value.to_bool orient in + let repeat = to_multi repeat in + let c = uthaw constr_with_bindings c in + (orient, repeat, c) +| _ -> assert false + +let rewriting = make_to_repr to_rewriting + +let to_debug v = match Value.to_int v with +| 0 -> Hints.Off +| 1 -> Hints.Info +| 2 -> Hints.Debug +| _ -> assert false + +let debug = make_to_repr to_debug + +let to_strategy v = match Value.to_int v with +| 0 -> Class_tactics.Bfs +| 1 -> Class_tactics.Dfs +| _ -> assert false + +let strategy = make_to_repr to_strategy + +let to_inversion_kind v = match Value.to_int v with +| 0 -> Inv.SimpleInversion +| 1 -> Inv.FullInversion +| 2 -> Inv.FullInversionClear +| _ -> assert false + +let inversion_kind = make_to_repr to_inversion_kind + +let to_move_location = function +| ValInt 0 -> Logic.MoveFirst +| ValInt 1 -> Logic.MoveLast +| ValBlk (0, [|id|]) -> Logic.MoveAfter (Value.to_ident id) +| ValBlk (1, [|id|]) -> Logic.MoveBefore (Value.to_ident id) +| _ -> assert false + +let move_location = make_to_repr to_move_location + +let to_generalize_arg v = match Value.to_tuple v with +| [| c; occs; na |] -> + (Value.to_constr c, to_occurrences occs, to_name na) +| _ -> assert false + +let generalize_arg = make_to_repr to_generalize_arg + +(** Standard tactics sharing their implementation with Ltac1 *) + +let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } + +let lift tac = tac <*> return v_unit + +let define_prim0 name tac = + let tac _ = lift tac in + Tac2env.define_primitive (pname name) (mk_closure arity_one tac) + +let define_prim1 name r0 f = + let tac x = lift (f (Value.repr_to r0 x)) in + Tac2env.define_primitive (pname name) (mk_closure arity_one tac) + +let define_prim2 name r0 r1 f = + let tac x y = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y)) in + Tac2env.define_primitive (pname name) (mk_closure (arity_suc arity_one) tac) + +let define_prim3 name r0 r1 r2 f = + let tac x y z = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y) (Value.repr_to r2 z)) in + Tac2env.define_primitive (pname name) (mk_closure (arity_suc (arity_suc arity_one)) tac) + +let define_prim4 name r0 r1 r2 r3 f = + let tac x y z u = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y) (Value.repr_to r2 z) (Value.repr_to r3 u)) in + Tac2env.define_primitive (pname name) (mk_closure (arity_suc (arity_suc (arity_suc arity_one))) tac) + +let define_prim5 name r0 r1 r2 r3 r4 f = + let tac x y z u v = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y) (Value.repr_to r2 z) (Value.repr_to r3 u) (Value.repr_to r4 v)) in + Tac2env.define_primitive (pname name) (mk_closure (arity_suc (arity_suc (arity_suc (arity_suc arity_one)))) tac) + +(** Tactics from Tacexpr *) + +let () = define_prim2 "tac_intros" bool intro_patterns begin fun ev ipat -> + Tac2tactics.intros_patterns ev ipat +end + +let () = define_prim4 "tac_apply" bool bool (list (thunk constr_with_bindings)) (option (pair ident (option intro_pattern))) begin fun adv ev cb ipat -> + Tac2tactics.apply adv ev cb ipat +end + +let () = define_prim3 "tac_elim" bool constr_with_bindings (option constr_with_bindings) begin fun ev c copt -> + Tac2tactics.elim ev c copt +end + +let () = define_prim2 "tac_case" bool constr_with_bindings begin fun ev c -> + Tac2tactics.general_case_analysis ev c +end + +let () = define_prim1 "tac_generalize" (list generalize_arg) begin fun cl -> + Tac2tactics.generalize cl +end + +let () = define_prim1 "tac_assert" assertion begin fun ast -> + Tac2tactics.assert_ ast +end + +let () = define_prim3 "tac_enough" constr (option (option (thunk unit))) (option intro_pattern) begin fun c tac ipat -> + let tac = Option.map (fun o -> Option.map (fun f -> thaw unit f) o) tac in + Tac2tactics.forward false tac ipat c +end + +let () = define_prim2 "tac_pose" name constr begin fun na c -> + Tactics.letin_tac None na c None Locusops.nowhere +end + +let () = define_prim3 "tac_set" bool (thunk (pair name constr)) clause begin fun ev p cl -> + Proofview.tclEVARMAP >>= fun sigma -> + thaw (pair name constr) p >>= fun (na, c) -> + Tac2tactics.letin_pat_tac ev None na (sigma, c) cl +end + +let () = define_prim5 "tac_remember" bool name (thunk constr) (option intro_pattern) clause begin fun ev na c eqpat cl -> + let eqpat = Option.default (IntroNaming IntroAnonymous) eqpat in + match eqpat with + | IntroNaming eqpat -> + Proofview.tclEVARMAP >>= fun sigma -> + thaw constr c >>= fun c -> + Tac2tactics.letin_pat_tac ev (Some (true, eqpat)) na (sigma, c) cl + | _ -> + Tacticals.New.tclZEROMSG (Pp.str "Invalid pattern for remember") +end + +let () = define_prim3 "tac_destruct" bool (list induction_clause) (option constr_with_bindings) begin fun ev ic using -> + Tac2tactics.induction_destruct false ev ic using +end + +let () = define_prim3 "tac_induction" bool (list induction_clause) (option constr_with_bindings) begin fun ev ic using -> + Tac2tactics.induction_destruct true ev ic using +end + +let () = define_prim1 "tac_red" clause begin fun cl -> + Tac2tactics.reduce (Red false) cl +end + +let () = define_prim1 "tac_hnf" clause begin fun cl -> + Tac2tactics.reduce Hnf cl +end + +let () = define_prim3 "tac_simpl" red_flags (option pattern_with_occs) clause begin fun flags where cl -> + Tac2tactics.simpl flags where cl +end + +let () = define_prim2 "tac_cbv" red_flags clause begin fun flags cl -> + Tac2tactics.cbv flags cl +end + +let () = define_prim2 "tac_cbn" red_flags clause begin fun flags cl -> + Tac2tactics.cbn flags cl +end + +let () = define_prim2 "tac_lazy" red_flags clause begin fun flags cl -> + Tac2tactics.lazy_ flags cl +end + +let () = define_prim2 "tac_unfold" (list reference_with_occs) clause begin fun refs cl -> + Tac2tactics.unfold refs cl +end + +let () = define_prim2 "tac_fold" (list constr) clause begin fun args cl -> + Tac2tactics.reduce (Fold args) cl +end + +let () = define_prim2 "tac_pattern" (list constr_with_occs) clause begin fun where cl -> + Tac2tactics.pattern where cl +end + +let () = define_prim2 "tac_vm" (option pattern_with_occs) clause begin fun where cl -> + Tac2tactics.vm where cl +end + +let () = define_prim2 "tac_native" (option pattern_with_occs) clause begin fun where cl -> + Tac2tactics.native where cl +end + +(** Reduction functions *) + +let lift tac = tac >>= fun c -> Proofview.tclUNIT (Value.of_constr c) + +let define_red1 name r0 f = + let tac x = lift (f (Value.repr_to r0 x)) in + Tac2env.define_primitive (pname name) (mk_closure arity_one tac) + +let define_red2 name r0 r1 f = + let tac x y = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y)) in + Tac2env.define_primitive (pname name) (mk_closure (arity_suc arity_one) tac) + +let define_red3 name r0 r1 r2 f = + let tac x y z = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y) (Value.repr_to r2 z)) in + Tac2env.define_primitive (pname name) (mk_closure (arity_suc (arity_suc arity_one)) tac) + +let () = define_red1 "eval_red" constr begin fun c -> + Tac2tactics.eval_red c +end + +let () = define_red1 "eval_hnf" constr begin fun c -> + Tac2tactics.eval_hnf c +end + +let () = define_red3 "eval_simpl" red_flags (option pattern_with_occs) constr begin fun flags where c -> + Tac2tactics.eval_simpl flags where c +end + +let () = define_red2 "eval_cbv" red_flags constr begin fun flags c -> + Tac2tactics.eval_cbv flags c +end + +let () = define_red2 "eval_cbn" red_flags constr begin fun flags c -> + Tac2tactics.eval_cbn flags c +end + +let () = define_red2 "eval_lazy" red_flags constr begin fun flags c -> + Tac2tactics.eval_lazy flags c +end + +let () = define_red2 "eval_unfold" (list reference_with_occs) constr begin fun refs c -> + Tac2tactics.eval_unfold refs c +end + +let () = define_red2 "eval_fold" (list constr) constr begin fun args c -> + Tac2tactics.eval_fold args c +end + +let () = define_red2 "eval_pattern" (list constr_with_occs) constr begin fun where c -> + Tac2tactics.eval_pattern where c +end + +let () = define_red2 "eval_vm" (option pattern_with_occs) constr begin fun where c -> + Tac2tactics.eval_vm where c +end + +let () = define_red2 "eval_native" (option pattern_with_occs) constr begin fun where c -> + Tac2tactics.eval_native where c +end + +let () = define_prim3 "tac_change" (option pattern) (fun1 (array constr) constr) clause begin fun pat c cl -> + Tac2tactics.change pat c cl +end + +let () = define_prim4 "tac_rewrite" bool (list rewriting) clause (option (thunk unit)) begin fun ev rw cl by -> + Tac2tactics.rewrite ev rw cl by +end + +let () = define_prim4 "tac_inversion" inversion_kind destruction_arg (option intro_pattern) (option (list ident)) begin fun knd arg pat ids -> + Tac2tactics.inversion knd arg pat ids +end + +(** Tactics from coretactics *) + +let () = define_prim0 "tac_reflexivity" Tactics.intros_reflexivity + +let () = define_prim2 "tac_move" ident move_location begin fun id mv -> + Tactics.move_hyp id mv +end + +let () = define_prim2 "tac_intro" (option ident) (option move_location) begin fun id mv -> + let mv = Option.default Logic.MoveLast mv in + Tactics.intro_move id mv +end + +(* + +TACTIC EXTEND exact + [ "exact" casted_constr(c) ] -> [ Tactics.exact_no_check c ] +END + +*) + +let () = define_prim0 "tac_assumption" Tactics.assumption + +let () = define_prim1 "tac_transitivity" constr begin fun c -> + Tactics.intros_transitivity (Some c) +end + +let () = define_prim0 "tac_etransitivity" (Tactics.intros_transitivity None) + +let () = define_prim1 "tac_cut" constr begin fun c -> + Tactics.cut c +end + +let () = define_prim2 "tac_left" bool bindings begin fun ev bnd -> + Tac2tactics.left_with_bindings ev bnd +end +let () = define_prim2 "tac_right" bool bindings begin fun ev bnd -> + Tac2tactics.right_with_bindings ev bnd +end + +let () = define_prim1 "tac_introsuntil" qhyp begin fun h -> + Tactics.intros_until h +end + +let () = define_prim1 "tac_exactnocheck" constr begin fun c -> + Tactics.exact_no_check c +end + +let () = define_prim1 "tac_vmcastnocheck" constr begin fun c -> + Tactics.vm_cast_no_check c +end + +let () = define_prim1 "tac_nativecastnocheck" constr begin fun c -> + Tactics.native_cast_no_check c +end + +let () = define_prim1 "tac_constructor" bool begin fun ev -> + Tactics.any_constructor ev None +end + +let () = define_prim3 "tac_constructorn" bool int bindings begin fun ev n bnd -> + Tac2tactics.constructor_tac ev None n bnd +end + +let () = define_prim2 "tac_specialize" constr_with_bindings (option intro_pattern) begin fun c ipat -> + Tac2tactics.specialize c ipat +end + +let () = define_prim1 "tac_symmetry" clause begin fun cl -> + Tac2tactics.symmetry cl +end + +let () = define_prim2 "tac_split" bool bindings begin fun ev bnd -> + Tac2tactics.split_with_bindings ev bnd +end + +let () = define_prim1 "tac_rename" (list (pair ident ident)) begin fun ids -> + Tactics.rename_hyp ids +end + +let () = define_prim1 "tac_revert" (list ident) begin fun ids -> + Tactics.revert ids +end + +let () = define_prim0 "tac_admit" Proofview.give_up + +let () = define_prim2 "tac_fix" ident int begin fun ident n -> + Tactics.fix ident n +end + +let () = define_prim1 "tac_cofix" ident begin fun ident -> + Tactics.cofix ident +end + +let () = define_prim1 "tac_clear" (list ident) begin fun ids -> + Tactics.clear ids +end + +let () = define_prim1 "tac_keep" (list ident) begin fun ids -> + Tactics.keep ids +end + +let () = define_prim1 "tac_clearbody" (list ident) begin fun ids -> + Tactics.clear_body ids +end + +(** Tactics from extratactics *) + +let () = define_prim2 "tac_discriminate" bool (option destruction_arg) begin fun ev arg -> + Tac2tactics.discriminate ev arg +end + +let () = define_prim3 "tac_injection" bool (option intro_patterns) (option destruction_arg) begin fun ev ipat arg -> + Tac2tactics.injection ev ipat arg +end + +let () = define_prim1 "tac_absurd" constr begin fun c -> + Contradiction.absurd c +end + +let () = define_prim1 "tac_contradiction" (option constr_with_bindings) begin fun c -> + Tac2tactics.contradiction c +end + +let () = define_prim4 "tac_autorewrite" bool (option (thunk unit)) (list ident) clause begin fun all by ids cl -> + Tac2tactics.autorewrite ~all by ids cl +end + +let () = define_prim1 "tac_subst" (list ident) begin fun ids -> + Equality.subst ids +end + +let () = define_prim0 "tac_substall" (return () >>= fun () -> Equality.subst_all ()) + +(** Auto *) + +let () = define_prim3 "tac_trivial" debug (list (thunk constr)) (option (list ident)) begin fun dbg lems dbs -> + Tac2tactics.trivial dbg lems dbs +end + +let () = define_prim5 "tac_eauto" debug (option int) (option int) (list (thunk constr)) (option (list ident)) begin fun dbg n p lems dbs -> + Tac2tactics.eauto dbg n p lems dbs +end + +let () = define_prim4 "tac_auto" debug (option int) (list (thunk constr)) (option (list ident)) begin fun dbg n lems dbs -> + Tac2tactics.auto dbg n lems dbs +end + +let () = define_prim4 "tac_newauto" debug (option int) (list (thunk constr)) (option (list ident)) begin fun dbg n lems dbs -> + Tac2tactics.new_auto dbg n lems dbs +end + +let () = define_prim3 "tac_typeclasses_eauto" (option strategy) (option int) (option (list ident)) begin fun str n dbs -> + Tac2tactics.typeclasses_eauto str n dbs +end + +(** Firstorder *) + +let () = define_prim3 "tac_firstorder" (option (thunk unit)) (list reference) (list ident) begin fun tac refs ids -> + Tac2tactics.firstorder tac refs ids +end diff --git a/vendor/Ltac2/src/tac2stdlib.mli b/vendor/Ltac2/src/tac2stdlib.mli new file mode 100644 index 0000000000..927b57074d --- /dev/null +++ b/vendor/Ltac2/src/tac2stdlib.mli @@ -0,0 +1,9 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Tactypes.ImplicitBindings l +| ExplicitBindings l -> + let l = List.map CAst.make l in + Tactypes.ExplicitBindings l +| NoBindings -> Tactypes.NoBindings + +let mk_with_bindings (x, b) = (x, mk_bindings b) + +let rec mk_intro_pattern = function +| IntroForthcoming b -> CAst.make @@ Tactypes.IntroForthcoming b +| IntroNaming ipat -> CAst.make @@ Tactypes.IntroNaming (mk_intro_pattern_naming ipat) +| IntroAction ipat -> CAst.make @@ Tactypes.IntroAction (mk_intro_pattern_action ipat) + +and mk_intro_pattern_naming = function +| IntroIdentifier id -> Namegen.IntroIdentifier id +| IntroFresh id -> Namegen.IntroFresh id +| IntroAnonymous -> Namegen.IntroAnonymous + +and mk_intro_pattern_action = function +| IntroWildcard -> Tactypes.IntroWildcard +| IntroOrAndPattern ipat -> Tactypes.IntroOrAndPattern (mk_or_and_intro_pattern ipat) +| IntroInjection ipats -> Tactypes.IntroInjection (List.map mk_intro_pattern ipats) +| IntroApplyOn (c, ipat) -> + let c = CAst.make @@ delayed_of_thunk Tac2ffi.constr c in + Tactypes.IntroApplyOn (c, mk_intro_pattern ipat) +| IntroRewrite b -> Tactypes.IntroRewrite b + +and mk_or_and_intro_pattern = function +| IntroOrPattern ipatss -> + Tactypes.IntroOrPattern (List.map (fun ipat -> List.map mk_intro_pattern ipat) ipatss) +| IntroAndPattern ipats -> + Tactypes.IntroAndPattern (List.map mk_intro_pattern ipats) + +let mk_intro_patterns ipat = List.map mk_intro_pattern ipat + +let mk_occurrences f = function +| AllOccurrences -> Locus.AllOccurrences +| AllOccurrencesBut l -> Locus.AllOccurrencesBut (List.map f l) +| NoOccurrences -> Locus.NoOccurrences +| OnlyOccurrences l -> Locus.OnlyOccurrences (List.map f l) + +let mk_occurrences_expr occ = + mk_occurrences (fun i -> Locus.ArgArg i) occ + +let mk_hyp_location (id, occs, h) = + ((mk_occurrences_expr occs, id), h) + +let mk_clause cl = { + Locus.onhyps = Option.map (fun l -> List.map mk_hyp_location l) cl.onhyps; + Locus.concl_occs = mk_occurrences_expr cl.concl_occs; +} + +let intros_patterns ev ipat = + let ipat = mk_intro_patterns ipat in + Tactics.intros_patterns ev ipat + +let apply adv ev cb cl = + let map c = + let c = thaw constr_with_bindings c >>= fun p -> return (mk_with_bindings p) in + None, CAst.make (delayed_of_tactic c) + in + let cb = List.map map cb in + match cl with + | None -> Tactics.apply_with_delayed_bindings_gen adv ev cb + | Some (id, cl) -> + let cl = Option.map mk_intro_pattern cl in + Tactics.apply_delayed_in adv ev id cb cl + +let mk_destruction_arg = function +| ElimOnConstr c -> + let c = c >>= fun c -> return (mk_with_bindings c) in + Tactics.ElimOnConstr (delayed_of_tactic c) +| ElimOnIdent id -> Tactics.ElimOnIdent CAst.(make id) +| ElimOnAnonHyp n -> Tactics.ElimOnAnonHyp n + +let mk_induction_clause (arg, eqn, as_, occ) = + let eqn = Option.map (fun ipat -> CAst.make @@ mk_intro_pattern_naming ipat) eqn in + let as_ = Option.map (fun ipat -> CAst.make @@ mk_or_and_intro_pattern ipat) as_ in + let occ = Option.map mk_clause occ in + ((None, mk_destruction_arg arg), (eqn, as_), occ) + +let induction_destruct isrec ev (ic : induction_clause list) using = + let ic = List.map mk_induction_clause ic in + let using = Option.map mk_with_bindings using in + Tactics.induction_destruct isrec ev (ic, using) + +let elim ev c copt = + let c = mk_with_bindings c in + let copt = Option.map mk_with_bindings copt in + Tactics.elim ev None c copt + +let generalize pl = + let mk_occ occs = mk_occurrences (fun i -> i) occs in + let pl = List.map (fun (c, occs, na) -> (mk_occ occs, c), na) pl in + Tactics.new_generalize_gen pl + +let general_case_analysis ev c = + let c = mk_with_bindings c in + Tactics.general_case_analysis ev None c + +let constructor_tac ev n i bnd = + let bnd = mk_bindings bnd in + Tactics.constructor_tac ev n i bnd + +let left_with_bindings ev bnd = + let bnd = mk_bindings bnd in + Tactics.left_with_bindings ev bnd + +let right_with_bindings ev bnd = + let bnd = mk_bindings bnd in + Tactics.right_with_bindings ev bnd + +let split_with_bindings ev bnd = + let bnd = mk_bindings bnd in + Tactics.split_with_bindings ev [bnd] + +let specialize c pat = + let c = mk_with_bindings c in + let pat = Option.map mk_intro_pattern pat in + Tactics.specialize c pat + +let change pat c cl = + let open Tac2ffi in + Proofview.Goal.enter begin fun gl -> + let c subst env sigma = + let subst = Array.map_of_list snd (Id.Map.bindings subst) in + delayed_of_tactic (Tac2ffi.app_fun1 c (array constr) constr subst) env sigma + in + let cl = mk_clause cl in + Tactics.change pat c cl + end + +let rewrite ev rw cl by = + let map_rw (orient, repeat, c) = + let c = c >>= fun c -> return (mk_with_bindings c) in + (Option.default true orient, repeat, None, delayed_of_tactic c) + in + let rw = List.map map_rw rw in + let cl = mk_clause cl in + let by = Option.map (fun tac -> Tacticals.New.tclCOMPLETE (thaw Tac2ffi.unit tac), Equality.Naive) by in + Equality.general_multi_rewrite ev rw cl by + +let symmetry cl = + let cl = mk_clause cl in + Tactics.intros_symmetry cl + +let forward fst tac ipat c = + let ipat = Option.map mk_intro_pattern ipat in + Tactics.forward fst tac ipat c + +let assert_ = function +| AssertValue (id, c) -> + let ipat = CAst.make @@ Tactypes.IntroNaming (Namegen.IntroIdentifier id) in + Tactics.forward true None (Some ipat) c +| AssertType (ipat, c, tac) -> + let ipat = Option.map mk_intro_pattern ipat in + let tac = Option.map (fun tac -> thaw Tac2ffi.unit tac) tac in + Tactics.forward true (Some tac) ipat c + +let letin_pat_tac ev ipat na c cl = + let ipat = Option.map (fun (b, ipat) -> (b, CAst.make @@ mk_intro_pattern_naming ipat)) ipat in + let cl = mk_clause cl in + Tactics.letin_pat_tac ev ipat na c cl + +(** Ltac interface treats differently global references than other term + arguments in reduction expressions. In Ltac1, this is done at parsing time. + Instead, we parse indifferently any pattern and dispatch when the tactic is + called. *) +let map_pattern_with_occs (pat, occ) = match pat with +| Pattern.PRef (ConstRef cst) -> (mk_occurrences_expr occ, Inl (EvalConstRef cst)) +| Pattern.PRef (VarRef id) -> (mk_occurrences_expr occ, Inl (EvalVarRef id)) +| _ -> (mk_occurrences_expr occ, Inr pat) + +let get_evaluable_reference = function +| VarRef id -> Proofview.tclUNIT (EvalVarRef id) +| ConstRef cst -> Proofview.tclUNIT (EvalConstRef cst) +| r -> + Tacticals.New.tclZEROMSG (str "Cannot coerce" ++ spc () ++ + Nametab.pr_global_env Id.Set.empty r ++ spc () ++ + str "to an evaluable reference.") + +let reduce r cl = + let cl = mk_clause cl in + Tactics.reduce r cl + +let simpl flags where cl = + let where = Option.map map_pattern_with_occs where in + let cl = mk_clause cl in + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + Tactics.reduce (Simpl (flags, where)) cl + +let cbv flags cl = + let cl = mk_clause cl in + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + Tactics.reduce (Cbv flags) cl + +let cbn flags cl = + let cl = mk_clause cl in + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + Tactics.reduce (Cbn flags) cl + +let lazy_ flags cl = + let cl = mk_clause cl in + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + Tactics.reduce (Lazy flags) cl + +let unfold occs cl = + let cl = mk_clause cl in + let map (gr, occ) = + let occ = mk_occurrences_expr occ in + get_evaluable_reference gr >>= fun gr -> Proofview.tclUNIT (occ, gr) + in + Proofview.Monad.List.map map occs >>= fun occs -> + Tactics.reduce (Unfold occs) cl + +let pattern where cl = + let where = List.map (fun (c, occ) -> (mk_occurrences_expr occ, c)) where in + let cl = mk_clause cl in + Tactics.reduce (Pattern where) cl + +let vm where cl = + let where = Option.map map_pattern_with_occs where in + let cl = mk_clause cl in + Tactics.reduce (CbvVm where) cl + +let native where cl = + let where = Option.map map_pattern_with_occs where in + let cl = mk_clause cl in + Tactics.reduce (CbvNative where) cl + +let eval_fun red c = + Tac2core.pf_apply begin fun env sigma -> + let (redfun, _) = Redexpr.reduction_of_red_expr env red in + let (sigma, ans) = redfun env sigma c in + Proofview.Unsafe.tclEVARS sigma >>= fun () -> + Proofview.tclUNIT ans + end + +let eval_red c = + eval_fun (Red false) c + +let eval_hnf c = + eval_fun Hnf c + +let eval_simpl flags where c = + let where = Option.map map_pattern_with_occs where in + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + eval_fun (Simpl (flags, where)) c + +let eval_cbv flags c = + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + eval_fun (Cbv flags) c + +let eval_cbn flags c = + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + eval_fun (Cbn flags) c + +let eval_lazy flags c = + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + eval_fun (Lazy flags) c + +let eval_unfold occs c = + let map (gr, occ) = + let occ = mk_occurrences_expr occ in + get_evaluable_reference gr >>= fun gr -> Proofview.tclUNIT (occ, gr) + in + Proofview.Monad.List.map map occs >>= fun occs -> + eval_fun (Unfold occs) c + +let eval_fold cl c = + eval_fun (Fold cl) c + +let eval_pattern where c = + let where = List.map (fun (pat, occ) -> (mk_occurrences_expr occ, pat)) where in + eval_fun (Pattern where) c + +let eval_vm where c = + let where = Option.map map_pattern_with_occs where in + eval_fun (CbvVm where) c + +let eval_native where c = + let where = Option.map map_pattern_with_occs where in + eval_fun (CbvNative where) c + +let on_destruction_arg tac ev arg = + Proofview.Goal.enter begin fun gl -> + match arg with + | None -> tac ev None + | Some (clear, arg) -> + let arg = match arg with + | ElimOnConstr c -> + let env = Proofview.Goal.env gl in + Proofview.tclEVARMAP >>= fun sigma -> + c >>= fun (c, lbind) -> + let lbind = mk_bindings lbind in + Proofview.tclEVARMAP >>= fun sigma' -> + let flags = tactic_infer_flags ev in + let (sigma', c) = Unification.finish_evar_resolution ~flags env sigma' (sigma, c) in + Proofview.tclUNIT (Some sigma', Tactics.ElimOnConstr (c, lbind)) + | ElimOnIdent id -> Proofview.tclUNIT (None, Tactics.ElimOnIdent CAst.(make id)) + | ElimOnAnonHyp n -> Proofview.tclUNIT (None, Tactics.ElimOnAnonHyp n) + in + arg >>= fun (sigma', arg) -> + let arg = Some (clear, arg) in + match sigma' with + | None -> tac ev arg + | Some sigma' -> + Tacticals.New.tclWITHHOLES ev (tac ev arg) sigma' + end + +let discriminate ev arg = + let arg = Option.map (fun arg -> None, arg) arg in + on_destruction_arg Equality.discr_tac ev arg + +let injection ev ipat arg = + let arg = Option.map (fun arg -> None, arg) arg in + let ipat = Option.map mk_intro_patterns ipat in + let tac ev arg = Equality.injClause None ipat ev arg in + on_destruction_arg tac ev arg + +let autorewrite ~all by ids cl = + let conds = if all then Some Equality.AllMatches else None in + let ids = List.map Id.to_string ids in + let cl = mk_clause cl in + match by with + | None -> Autorewrite.auto_multi_rewrite ?conds ids cl + | Some by -> + let by = thaw Tac2ffi.unit by in + Autorewrite.auto_multi_rewrite_with ?conds by ids cl + +(** Auto *) + +let trivial debug lems dbs = + let lems = List.map (fun c -> delayed_of_thunk Tac2ffi.constr c) lems in + let dbs = Option.map (fun l -> List.map Id.to_string l) dbs in + Auto.h_trivial ~debug lems dbs + +let auto debug n lems dbs = + let lems = List.map (fun c -> delayed_of_thunk Tac2ffi.constr c) lems in + let dbs = Option.map (fun l -> List.map Id.to_string l) dbs in + Auto.h_auto ~debug n lems dbs + +let new_auto debug n lems dbs = + let make_depth n = snd (Eauto.make_dimension n None) in + let lems = List.map (fun c -> delayed_of_thunk Tac2ffi.constr c) lems in + match dbs with + | None -> Auto.new_full_auto ~debug (make_depth n) lems + | Some dbs -> + let dbs = List.map Id.to_string dbs in + Auto.new_auto ~debug (make_depth n) lems dbs + +let eauto debug n p lems dbs = + let lems = List.map (fun c -> delayed_of_thunk Tac2ffi.constr c) lems in + let dbs = Option.map (fun l -> List.map Id.to_string l) dbs in + Eauto.gen_eauto (Eauto.make_dimension n p) lems dbs + +let typeclasses_eauto strategy depth dbs = + let only_classes, dbs = match dbs with + | None -> + true, [Class_tactics.typeclasses_db] + | Some dbs -> + let dbs = List.map Id.to_string dbs in + false, dbs + in + Class_tactics.typeclasses_eauto ~only_classes ?strategy ~depth dbs + +(** Inversion *) + +let inversion knd arg pat ids = + let ids = match ids with + | None -> [] + | Some l -> l + in + begin match pat with + | None -> Proofview.tclUNIT None + | Some (IntroAction (IntroOrAndPattern p)) -> + Proofview.tclUNIT (Some (CAst.make @@ mk_or_and_intro_pattern p)) + | Some _ -> + Tacticals.New.tclZEROMSG (str "Inversion only accept disjunctive patterns") + end >>= fun pat -> + let inversion _ arg = + begin match arg with + | None -> assert false + | Some (_, Tactics.ElimOnAnonHyp n) -> + Inv.inv_clause knd pat ids (AnonHyp n) + | Some (_, Tactics.ElimOnIdent {CAst.v=id}) -> + Inv.inv_clause knd pat ids (NamedHyp id) + | Some (_, Tactics.ElimOnConstr c) -> + let open Tactypes in + let anon = CAst.make @@ IntroNaming Namegen.IntroAnonymous in + Tactics.specialize c (Some anon) >>= fun () -> + Tacticals.New.onLastHypId (fun id -> Inv.inv_clause knd pat ids (NamedHyp id)) + end + in + on_destruction_arg inversion true (Some (None, arg)) + +let contradiction c = + let c = Option.map mk_with_bindings c in + Contradiction.contradiction c + +(** Firstorder *) + +let firstorder tac refs ids = + let open Ground_plugin in + let ids = List.map Id.to_string ids in + let tac = Option.map (fun tac -> thaw Tac2ffi.unit tac) tac in + G_ground.gen_ground_tac true tac refs ids diff --git a/vendor/Ltac2/src/tac2tactics.mli b/vendor/Ltac2/src/tac2tactics.mli new file mode 100644 index 0000000000..026673acbf --- /dev/null +++ b/vendor/Ltac2/src/tac2tactics.mli @@ -0,0 +1,124 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* intro_pattern list -> unit tactic + +val apply : advanced_flag -> evars_flag -> + constr_with_bindings thunk list -> + (Id.t * intro_pattern option) option -> unit tactic + +val induction_destruct : rec_flag -> evars_flag -> + induction_clause list -> constr_with_bindings option -> unit tactic + +val elim : evars_flag -> constr_with_bindings -> constr_with_bindings option -> + unit tactic + +val general_case_analysis : evars_flag -> constr_with_bindings -> unit tactic + +val generalize : (constr * occurrences * Name.t) list -> unit tactic + +val constructor_tac : evars_flag -> int option -> int -> bindings -> unit tactic + +val left_with_bindings : evars_flag -> bindings -> unit tactic +val right_with_bindings : evars_flag -> bindings -> unit tactic +val split_with_bindings : evars_flag -> bindings -> unit tactic + +val specialize : constr_with_bindings -> intro_pattern option -> unit tactic + +val change : Pattern.constr_pattern option -> (constr array, constr) Tac2ffi.fun1 -> clause -> unit tactic + +val rewrite : + evars_flag -> rewriting list -> clause -> unit thunk option -> unit tactic + +val symmetry : clause -> unit tactic + +val forward : bool -> unit tactic option option -> + intro_pattern option -> constr -> unit tactic + +val assert_ : assertion -> unit tactic + +val letin_pat_tac : evars_flag -> (bool * intro_pattern_naming) option -> + Name.t -> (Evd.evar_map * constr) -> clause -> unit tactic + +val reduce : Redexpr.red_expr -> clause -> unit tactic + +val simpl : GlobRef.t glob_red_flag -> + (Pattern.constr_pattern * occurrences) option -> clause -> unit tactic + +val cbv : GlobRef.t glob_red_flag -> clause -> unit tactic + +val cbn : GlobRef.t glob_red_flag -> clause -> unit tactic + +val lazy_ : GlobRef.t glob_red_flag -> clause -> unit tactic + +val unfold : (GlobRef.t * occurrences) list -> clause -> unit tactic + +val pattern : (constr * occurrences) list -> clause -> unit tactic + +val vm : (Pattern.constr_pattern * occurrences) option -> clause -> unit tactic + +val native : (Pattern.constr_pattern * occurrences) option -> clause -> unit tactic + +val eval_red : constr -> constr tactic + +val eval_hnf : constr -> constr tactic + +val eval_simpl : GlobRef.t glob_red_flag -> + (Pattern.constr_pattern * occurrences) option -> constr -> constr tactic + +val eval_cbv : GlobRef.t glob_red_flag -> constr -> constr tactic + +val eval_cbn : GlobRef.t glob_red_flag -> constr -> constr tactic + +val eval_lazy : GlobRef.t glob_red_flag -> constr -> constr tactic + +val eval_unfold : (GlobRef.t * occurrences) list -> constr -> constr tactic + +val eval_fold : constr list -> constr -> constr tactic + +val eval_pattern : (EConstr.t * occurrences) list -> constr -> constr tactic + +val eval_vm : (Pattern.constr_pattern * occurrences) option -> constr -> constr tactic + +val eval_native : (Pattern.constr_pattern * occurrences) option -> constr -> constr tactic + +val discriminate : evars_flag -> destruction_arg option -> unit tactic + +val injection : evars_flag -> intro_pattern list option -> destruction_arg option -> unit tactic + +val autorewrite : all:bool -> unit thunk option -> Id.t list -> clause -> unit tactic + +val trivial : Hints.debug -> constr thunk list -> Id.t list option -> + unit Proofview.tactic + +val auto : Hints.debug -> int option -> constr thunk list -> + Id.t list option -> unit Proofview.tactic + +val new_auto : Hints.debug -> int option -> constr thunk list -> + Id.t list option -> unit Proofview.tactic + +val eauto : Hints.debug -> int option -> int option -> constr thunk list -> + Id.t list option -> unit Proofview.tactic + +val typeclasses_eauto : Class_tactics.search_strategy option -> int option -> + Id.t list option -> unit Proofview.tactic + +val inversion : Inv.inversion_kind -> destruction_arg -> intro_pattern option -> Id.t list option -> unit tactic + +val contradiction : constr_with_bindings option -> unit tactic + +val firstorder : unit thunk option -> GlobRef.t list -> Id.t list -> unit tactic diff --git a/vendor/Ltac2/src/tac2types.mli b/vendor/Ltac2/src/tac2types.mli new file mode 100644 index 0000000000..fa31153a27 --- /dev/null +++ b/vendor/Ltac2/src/tac2types.mli @@ -0,0 +1,92 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* $@ + if [ $$? = 0 ]; then \ + echo " $<... OK"; \ + else \ + echo " $<... FAIL!"; \ + fi; \ + +clean: + rm -f *.log diff --git a/vendor/Ltac2/tests/compat.v b/vendor/Ltac2/tests/compat.v new file mode 100644 index 0000000000..489fa638e4 --- /dev/null +++ b/vendor/Ltac2/tests/compat.v @@ -0,0 +1,58 @@ +Require Import Ltac2.Ltac2. + +Import Ltac2.Notations. + +(** Test calls to Ltac1 from Ltac2 *) + +Ltac2 foo () := ltac1:(discriminate). + +Goal true = false -> False. +Proof. +foo (). +Qed. + +Goal true = false -> false = true. +Proof. +intros H; ltac1:(match goal with [ H : ?P |- _ ] => rewrite H end); reflexivity. +Qed. + +Goal true = false -> false = true. +Proof. +intros H; ltac1:(rewrite H); reflexivity. +Abort. + +(** Variables do not cross the compatibility layer boundary. *) +Fail Ltac2 bar nay := ltac1:(discriminate nay). + +Fail Ltac2 pose1 (v : constr) := + ltac1:(pose $v). + +(** Test calls to Ltac2 from Ltac1 *) + +Set Default Proof Mode "Classic". + +Ltac foo := ltac2:(foo ()). + +Goal true = false -> False. +Proof. +ltac2:(foo ()). +Qed. + +Goal true = false -> False. +Proof. +foo. +Qed. + +(** Variables do not cross the compatibility layer boundary. *) +Fail Ltac bar x := ltac2:(foo x). + +Ltac mytac tac := idtac "wow". + +Goal True. +Proof. +(** Fails because quotation is evaluated eagerly *) +Fail mytac ltac2:(fail). +(** One has to thunk thanks to the idtac trick *) +let t := idtac; ltac2:(fail) in mytac t. +constructor. +Qed. diff --git a/vendor/Ltac2/tests/errors.v b/vendor/Ltac2/tests/errors.v new file mode 100644 index 0000000000..c677f6af5d --- /dev/null +++ b/vendor/Ltac2/tests/errors.v @@ -0,0 +1,12 @@ +Require Import Ltac2.Ltac2. + +Goal True. +Proof. +let x := Control.plus + (fun () => let _ := constr:(nat -> 0) in 0) + (fun e => match e with Not_found => 1 | _ => 2 end) in +match Int.equal x 2 with +| true => () +| false => Control.throw (Tactic_failure None) +end. +Abort. diff --git a/vendor/Ltac2/tests/example1.v b/vendor/Ltac2/tests/example1.v new file mode 100644 index 0000000000..023791050f --- /dev/null +++ b/vendor/Ltac2/tests/example1.v @@ -0,0 +1,27 @@ +Require Import Ltac2.Ltac2. + +Import Ltac2.Control. + +(** Alternative implementation of the hyp primitive *) +Ltac2 get_hyp_by_name x := + let h := hyps () in + let rec find x l := match l with + | [] => zero Not_found + | p :: l => + match p with + | (id, _, t) => + match Ident.equal x id with + | true => t + | false => find x l + end + end + end in + find x h. + +Print Ltac2 get_hyp_by_name. + +Goal forall n m, n + m = 0 -> n = 0. +Proof. +refine (fun () => '(fun n m H => _)). +let t := get_hyp_by_name @H in Message.print (Message.of_constr t). +Abort. diff --git a/vendor/Ltac2/tests/example2.v b/vendor/Ltac2/tests/example2.v new file mode 100644 index 0000000000..c953d25061 --- /dev/null +++ b/vendor/Ltac2/tests/example2.v @@ -0,0 +1,281 @@ +Require Import Ltac2.Ltac2. + +Import Ltac2.Notations. + +Set Default Goal Selector "all". + +Goal exists n, n = 0. +Proof. +split with (x := 0). +reflexivity. +Qed. + +Goal exists n, n = 0. +Proof. +split with 0. +split. +Qed. + +Goal exists n, n = 0. +Proof. +let myvar := Std.NamedHyp @x in split with ($myvar := 0). +split. +Qed. + +Goal (forall n : nat, n = 0 -> False) -> True. +Proof. +intros H. +eelim &H. +split. +Qed. + +Goal (forall n : nat, n = 0 -> False) -> True. +Proof. +intros H. +elim &H with 0. +split. +Qed. + +Goal forall (P : nat -> Prop), (forall n m, n = m -> P n) -> P 0. +Proof. +intros P H. +Fail apply &H. +apply &H with (m := 0). +split. +Qed. + +Goal forall (P : nat -> Prop), (forall n m, n = m -> P n) -> (0 = 1) -> P 0. +Proof. +intros P H e. +apply &H with (m := 1) in e. +exact e. +Qed. + +Goal forall (P : nat -> Prop), (forall n m, n = m -> P n) -> P 0. +Proof. +intros P H. +eapply &H. +split. +Qed. + +Goal exists n, n = 0. +Proof. +Fail constructor 1. +constructor 1 with (x := 0). +split. +Qed. + +Goal exists n, n = 0. +Proof. +econstructor 1. +split. +Qed. + +Goal forall n, 0 + n = n. +Proof. +intros n. +induction &n as [|n] using nat_rect; split. +Qed. + +Goal forall n, 0 + n = n. +Proof. +intros n. +let n := @X in +let q := Std.NamedHyp @P in +induction &n as [|$n] using nat_rect with ($q := fun m => 0 + m = m); split. +Qed. + +Goal forall n, 0 + n = n. +Proof. +intros n. +destruct &n as [|n] using nat_rect; split. +Qed. + +Goal forall n, 0 + n = n. +Proof. +intros n. +let n := @X in +let q := Std.NamedHyp @P in +destruct &n as [|$n] using nat_rect with ($q := fun m => 0 + m = m); split. +Qed. + +Goal forall b1 b2, andb b1 b2 = andb b2 b1. +Proof. +intros b1 b2. +destruct &b1 as [|], &b2 as [|]; split. +Qed. + +Goal forall n m, n = 0 -> n + m = m. +Proof. +intros n m Hn. +rewrite &Hn; split. +Qed. + +Goal forall n m p, n = m -> p = m -> 0 = n -> p = 0. +Proof. +intros n m p He He' Hn. +rewrite &He, <- &He' in Hn. +rewrite &Hn. +split. +Qed. + +Goal forall n m, (m = n -> n = m) -> m = n -> n = 0 -> m = 0. +Proof. +intros n m He He' He''. +rewrite <- &He by assumption. +Control.refine (fun () => &He''). +Qed. + +Goal forall n (r := if true then n else 0), r = n. +Proof. +intros n r. +hnf in r. +split. +Qed. + +Goal 1 = 0 -> 0 = 0. +Proof. +intros H. +pattern 0 at 1. +let occ := 2 in pattern 1 at 1, 0 at $occ in H. +reflexivity. +Qed. + +Goal 1 + 1 = 2. +Proof. +vm_compute. +reflexivity. +Qed. + +Goal 1 + 1 = 2. +Proof. +native_compute. +reflexivity. +Qed. + +Goal 1 + 1 = 2 - 0 -> True. +Proof. +intros H. +vm_compute plus in H. +reflexivity. +Qed. + +Goal 1 = 0 -> True /\ True. +Proof. +intros H. +split; fold (1 + 0) (1 + 0) in H. +reflexivity. +Qed. + +Goal 1 + 1 = 2. +Proof. +cbv [ Nat.add ]. +reflexivity. +Qed. + +Goal 1 + 1 = 2. +Proof. +let x := reference:(Nat.add) in +cbn beta iota delta [ $x ]. +reflexivity. +Qed. + +Goal 1 + 1 = 2. +Proof. +simpl beta. +reflexivity. +Qed. + +Goal 1 + 1 = 2. +Proof. +lazy. +reflexivity. +Qed. + +Goal let x := 1 + 1 - 1 in x = x. +Proof. +intros x. +unfold &x at 1. +let x := reference:(Nat.sub) in unfold Nat.add, $x in x. +reflexivity. +Qed. + +Goal exists x y : nat, x = y. +Proof. +exists 0, 0; reflexivity. +Qed. + +Goal exists x y : nat, x = y. +Proof. +eexists _, 0; reflexivity. +Qed. + +Goal exists x y : nat, x = y. +Proof. +refine '(let x := 0 in _). +eexists; exists &x; reflexivity. +Qed. + +Goal True. +Proof. +pose (X := True). +constructor. +Qed. + +Goal True. +Proof. +pose True as X. +constructor. +Qed. + +Goal True. +Proof. +let x := @foo in +set ($x := True) in * |-. +constructor. +Qed. + +Goal 0 = 0. +Proof. +remember 0 as n eqn: foo at 1. +rewrite foo. +reflexivity. +Qed. + +Goal True. +Proof. +assert (H := 0 + 0). +constructor. +Qed. + +Goal True. +Proof. +assert (exists n, n = 0) as [n Hn]. ++ exists 0; reflexivity. ++ exact I. +Qed. + +Goal True -> True. +Proof. +assert (H : 0 + 0 = 0) by reflexivity. +intros x; exact x. +Qed. + +Goal 1 + 1 = 2. +Proof. +change (?a + 1 = 2) with (2 = $a + 1). +reflexivity. +Qed. + +Goal (forall n, n = 0 -> False) -> False. +Proof. +intros H. +specialize (H 0 eq_refl). +destruct H. +Qed. + +Goal (forall n, n = 0 -> False) -> False. +Proof. +intros H. +specialize (H 0 eq_refl) as []. +Qed. diff --git a/vendor/Ltac2/tests/matching.v b/vendor/Ltac2/tests/matching.v new file mode 100644 index 0000000000..4338cbd32f --- /dev/null +++ b/vendor/Ltac2/tests/matching.v @@ -0,0 +1,71 @@ +Require Import Ltac2.Ltac2 Ltac2.Notations. + +Ltac2 Type exn ::= [ Nope ]. + +Ltac2 check_id id id' := match Ident.equal id id' with +| true => () +| false => Control.throw Nope +end. + +Goal True -> False. +Proof. +Fail +let b := { contents := true } in +let f c := + match b.(contents) with + | true => Message.print (Message.of_constr c); b.(contents) := false; fail + | false => () + end +in +(** This fails because the matching is not allowed to backtrack once + it commits to a branch*) +lazy_match! '(nat -> bool) with context [?a] => f a end. +lazy_match! Control.goal () with ?a -> ?b => Message.print (Message.of_constr b) end. + +(** This one works by taking the second match context, i.e. ?a := nat *) +let b := { contents := true } in +let f c := + match b.(contents) with + | true => b.(contents) := false; fail + | false => Message.print (Message.of_constr c) + end +in +match! '(nat -> bool) with context [?a] => f a end. +Abort. + +Goal forall (i j : unit) (x y : nat) (b : bool), True. +Proof. +Fail match! goal with +| [ h : ?t, h' : ?t |- _ ] => () +end. +intros i j x y b. +match! goal with +| [ h : ?t, h' : ?t |- _ ] => + check_id h @x; + check_id h' @y +end. +match! reverse goal with +| [ h : ?t, h' : ?t |- _ ] => + check_id h @j; + check_id h' @i +end. +Abort. + +(* Check #79 *) +Goal 2 = 3. + Control.plus + (fun () + => lazy_match! goal with + | [ |- 2 = 3 ] => Control.zero (Tactic_failure None) + | [ |- 2 = _ ] => Control.zero (Tactic_failure (Some (Message.of_string "should not be printed"))) + end) + (fun e + => match e with + | Tactic_failure c + => match c with + | None => () + | _ => Control.zero e + end + | e => Control.zero e + end). +Abort. diff --git a/vendor/Ltac2/tests/quot.v b/vendor/Ltac2/tests/quot.v new file mode 100644 index 0000000000..624c4ad0c1 --- /dev/null +++ b/vendor/Ltac2/tests/quot.v @@ -0,0 +1,26 @@ +Require Import Ltac2.Ltac2. + +(** Test for quotations *) + +Ltac2 ref0 () := reference:(&x). +Ltac2 ref1 () := reference:(nat). +Ltac2 ref2 () := reference:(Datatypes.nat). +Fail Ltac2 ref () := reference:(i_certainly_dont_exist). +Fail Ltac2 ref () := reference:(And.Me.neither). + +Goal True. +Proof. +let x := constr:(I) in +let y := constr:((fun z => z) $x) in +Control.refine (fun _ => y). +Qed. + +Goal True. +Proof. +(** Here, Ltac2 should not put its variables in the same environment as + Ltac1 otherwise the second binding fails as x is bound but not an + ident. *) +let x := constr:(I) in +let y := constr:((fun x => x) $x) in +Control.refine (fun _ => y). +Qed. diff --git a/vendor/Ltac2/tests/rebind.v b/vendor/Ltac2/tests/rebind.v new file mode 100644 index 0000000000..e1c20a2059 --- /dev/null +++ b/vendor/Ltac2/tests/rebind.v @@ -0,0 +1,34 @@ +Require Import Ltac2.Ltac2 Ltac2.Notations. + +Ltac2 mutable foo () := constructor. + +Goal True. +Proof. +foo (). +Qed. + +Ltac2 Set foo := fun _ => fail. + +Goal True. +Proof. +Fail foo (). +constructor. +Qed. + +(** Not the right type *) +Fail Ltac2 Set foo := 0. + +Ltac2 bar () := (). + +(** Cannot redefine non-mutable tactics *) +Fail Ltac2 Set bar := fun _ => (). + +(** Subtype check *) + +Ltac2 mutable rec f x := f x. + +Fail Ltac2 Set f := fun x => x. + +Ltac2 mutable g x := x. + +Ltac2 Set g := f. diff --git a/vendor/Ltac2/tests/stuff/ltac2.v b/vendor/Ltac2/tests/stuff/ltac2.v new file mode 100644 index 0000000000..370bc70d15 --- /dev/null +++ b/vendor/Ltac2/tests/stuff/ltac2.v @@ -0,0 +1,143 @@ +Require Import Ltac2.Ltac2. + +Ltac2 foo (_ : int) := + let f (x : int) := x in + let _ := f 0 in + f 1. + +Print Ltac2 foo. + +Import Control. + +Ltac2 exact x := refine (fun () => x). + +Print Ltac2 refine. +Print Ltac2 exact. + +Ltac2 foo' () := ident:(bla). + +Print Ltac2 foo'. + +Ltac2 bar x h := match x with +| None => constr:(fun H => ltac2:(exact (hyp ident:(H))) -> nat) +| Some x => x +end. + +Print Ltac2 bar. + +Ltac2 qux := Some 0. + +Print Ltac2 qux. + +Ltac2 Type foo := [ Foo (int) ]. + +Fail Ltac2 qux0 := Foo None. + +Ltac2 Type 'a ref := { mutable contents : 'a }. + +Fail Ltac2 qux0 := { contents := None }. +Ltac2 foo0 () := { contents := None }. + +Print Ltac2 foo0. + +Ltac2 qux0 x := x.(contents). +Ltac2 qux1 x := x.(contents) := x.(contents). + +Ltac2 qux2 := ([1;2], true). + +Print Ltac2 qux0. +Print Ltac2 qux1. +Print Ltac2 qux2. + +Import Control. + +Ltac2 qux3 x := constr:(nat -> ltac2:(refine (fun () => hyp x))). + +Print Ltac2 qux3. + +Ltac2 Type rec nat := [ O | S (nat) ]. + +Ltac2 message_of_nat n := +let rec aux n := +match n with +| O => Message.of_string "O" +| S n => Message.concat (Message.of_string "S") (aux n) +end in aux n. + +Print Ltac2 message_of_nat. + +Ltac2 numgoals () := + let r := { contents := O } in + enter (fun () => r.(contents) := S (r.(contents))); + r.(contents). + +Print Ltac2 numgoals. + +Goal True /\ False. +Proof. +let n := numgoals () in Message.print (message_of_nat n). +refine (fun () => open_constr:((fun x => conj _ _) 0)); (). +let n := numgoals () in Message.print (message_of_nat n). + +Fail (hyp ident:(x)). +Fail (enter (fun () => hyp ident:(There_is_no_spoon); ())). + +enter (fun () => Message.print (Message.of_string "foo")). + +enter (fun () => Message.print (Message.of_constr (goal ()))). +Fail enter (fun () => Message.print (Message.of_constr (qux3 ident:(x)))). +enter (fun () => plus (fun () => constr:(_); ()) (fun _ => ())). +plus + (fun () => enter (fun () => let x := ident:(foo) in let _ := hyp x in ())) (fun _ => Message.print (Message.of_string "failed")). +let x := { contents := 0 } in +let x := x.(contents) := x.(contents) in x. +Abort. + +Ltac2 Type exn ::= [ Foo ]. + +Goal True. +Proof. +plus (fun () => zero Foo) (fun _ => ()). +Abort. + +Ltac2 Type exn ::= [ Bar (string) ]. + +Goal True. +Proof. +Fail zero (Bar "lol"). +Abort. + +Ltac2 Notation "refine!" c(thunk(constr)) := refine c. + +Goal True. +Proof. +refine! I. +Abort. + +Goal True. +Proof. +let x () := plus (fun () => 0) (fun _ => 1) in +match case x with +| Val x => + match x with + | (x, k) => Message.print (Message.of_int (k Not_found)) + end +| Err x => Message.print (Message.of_string "Err") +end. +Abort. + +Goal (forall n : nat, n = 0 -> False) -> True. +Proof. +refine (fun () => '(fun H => _)). +Std.case true (hyp @H, Std.ExplicitBindings [Std.NamedHyp @n, '0]). +refine (fun () => 'eq_refl). +Qed. + +Goal forall x, 1 + x = x + 1. +Proof. +refine (fun () => '(fun x => _)). +Std.cbv { + Std.rBeta := true; Std.rMatch := true; Std.rFix := true; Std.rCofix := true; + Std.rZeta := true; Std.rDelta := true; Std.rConst := []; +} { Std.on_hyps := None; Std.on_concl := Std.AllOccurrences }. +Abort. diff --git a/vendor/Ltac2/tests/tacticals.v b/vendor/Ltac2/tests/tacticals.v new file mode 100644 index 0000000000..1a2fbcbb37 --- /dev/null +++ b/vendor/Ltac2/tests/tacticals.v @@ -0,0 +1,34 @@ +Require Import Ltac2.Ltac2. + +Import Ltac2.Notations. + +Goal True. +Proof. +Fail fail. +Fail solve [ () ]. +try fail. +repeat fail. +repeat (). +solve [ constructor ]. +Qed. + +Goal True. +Proof. +first [ + Message.print (Message.of_string "Yay"); fail +| constructor +| Message.print (Message.of_string "I won't be printed") +]. +Qed. + +Goal True /\ True. +Proof. +Fail split > [ split | |]. +split > [split | split]. +Qed. + +Goal True /\ (True -> True) /\ True. +Proof. +split > [ | split] > [split | .. | split]. +intros H; refine &H. +Qed. diff --git a/vendor/Ltac2/tests/typing.v b/vendor/Ltac2/tests/typing.v new file mode 100644 index 0000000000..9f18292716 --- /dev/null +++ b/vendor/Ltac2/tests/typing.v @@ -0,0 +1,72 @@ +Require Import Ltac2.Ltac2. + +(** Ltac2 is typed à la ML. *) + +Ltac2 test0 n := Int.add n 1. + +Print Ltac2 test0. + +Ltac2 test1 () := test0 0. + +Print Ltac2 test1. + +Fail Ltac2 test2 () := test0 true. + +Fail Ltac2 test2 () := test0 0 0. + +Ltac2 test3 f x := x, (f x, x). + +Print Ltac2 test3. + +(** Polymorphism *) + +Ltac2 rec list_length l := +match l with +| [] => 0 +| x :: l => Int.add 1 (list_length l) +end. + +Print Ltac2 list_length. + +(** Pattern-matching *) + +Ltac2 ifb b f g := match b with +| true => f () +| false => g () +end. + +Print Ltac2 ifb. + +Ltac2 if_not_found e f g := match e with +| Not_found => f () +| _ => g () +end. + +Fail Ltac2 ifb' b f g := match b with +| true => f () +end. + +Fail Ltac2 if_not_found' e f g := match e with +| Not_found => f () +end. + +(** Reimplementing 'do'. Return value of the function useless. *) + +Ltac2 rec do n tac := match Int.equal n 0 with +| true => () +| false => tac (); do (Int.sub n 1) tac +end. + +Print Ltac2 do. + +(** Non-function pure values are OK. *) + +Ltac2 tuple0 := ([1; 2], true, (fun () => "yay")). + +Print Ltac2 tuple0. + +(** Impure values are not. *) + +Fail Ltac2 not_a_value := { contents := 0 }. +Fail Ltac2 not_a_value := "nope". +Fail Ltac2 not_a_value := list_length []. diff --git a/vendor/Ltac2/theories/Array.v b/vendor/Ltac2/theories/Array.v new file mode 100644 index 0000000000..11b64e3515 --- /dev/null +++ b/vendor/Ltac2/theories/Array.v @@ -0,0 +1,14 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* 'a -> 'a array := "ltac2" "array_make". +Ltac2 @external length : 'a array -> int := "ltac2" "array_length". +Ltac2 @external get : 'a array -> int -> 'a := "ltac2" "array_get". +Ltac2 @external set : 'a array -> int -> 'a -> unit := "ltac2" "array_set". diff --git a/vendor/Ltac2/theories/Char.v b/vendor/Ltac2/theories/Char.v new file mode 100644 index 0000000000..29fef60f2c --- /dev/null +++ b/vendor/Ltac2/theories/Char.v @@ -0,0 +1,12 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* char := "ltac2" "char_of_int". +Ltac2 @external to_int : char -> int := "ltac2" "char_to_int". diff --git a/vendor/Ltac2/theories/Constr.v b/vendor/Ltac2/theories/Constr.v new file mode 100644 index 0000000000..d8d222730e --- /dev/null +++ b/vendor/Ltac2/theories/Constr.v @@ -0,0 +1,72 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* constr := "ltac2" "constr_type". +(** Return the type of a term *) + +Ltac2 @ external equal : constr -> constr -> bool := "ltac2" "constr_equal". +(** Strict syntactic equality: only up to α-conversion and evar expansion *) + +Module Unsafe. + +(** Low-level access to kernel terms. Use with care! *) + +Ltac2 Type case. + +Ltac2 Type kind := [ +| Rel (int) +| Var (ident) +| Meta (meta) +| Evar (evar, constr array) +| Sort (sort) +| Cast (constr, cast, constr) +| Prod (ident option, constr, constr) +| Lambda (ident option, constr, constr) +| LetIn (ident option, constr, constr, constr) +| App (constr, constr array) +| Constant (constant, instance) +| Ind (inductive, instance) +| Constructor (constructor, instance) +| Case (case, constr, constr, constr array) +| Fix (int array, int, ident option array, constr array, constr array) +| CoFix (int, ident option array, constr array, constr array) +| Proj (projection, constr) +]. + +Ltac2 @ external kind : constr -> kind := "ltac2" "constr_kind". + +Ltac2 @ external make : kind -> constr := "ltac2" "constr_make". + +Ltac2 @ external check : constr -> constr result := "ltac2" "constr_check". +(** Checks that a constr generated by unsafe means is indeed safe in the + current environment, and returns it, or the error otherwise. Panics if + not focussed. *) + +Ltac2 @ external substnl : constr list -> int -> constr -> constr := "ltac2" "constr_substnl". +(** [substnl [r₁;...;rₙ] k c] substitutes in parallel [Rel(k+1); ...; Rel(k+n)] with + [r₁;...;rₙ] in [c]. *) + +Ltac2 @ external closenl : ident list -> int -> constr -> constr := "ltac2" "constr_closenl". +(** [closenl [x₁;...;xₙ] k c] abstracts over variables [x₁;...;xₙ] and replaces them with + [Rel(k); ...; Rel(k+n-1)] in [c]. If two names are identical, the one of least index is kept. *) + +Ltac2 @ external case : inductive -> case := "ltac2" "constr_case". +(** Generate the case information for a given inductive type. *) + +Ltac2 @ external constructor : inductive -> int -> constructor := "ltac2" "constr_constructor". +(** Generate the i-th constructor for a given inductive type. Indexing starts + at 0. Panics if there is no such constructor. *) + +End Unsafe. + +Ltac2 @ external in_context : ident -> constr -> (unit -> unit) -> constr := "ltac2" "constr_in_context". +(** On a focussed goal [Γ ⊢ A], [in_context id c tac] evaluates [tac] in a + focussed goal [Γ, id : c ⊢ ?X] and returns [fun (id : c) => t] where [t] is + the proof built by the tactic. *) diff --git a/vendor/Ltac2/theories/Control.v b/vendor/Ltac2/theories/Control.v new file mode 100644 index 0000000000..071c2ea8ce --- /dev/null +++ b/vendor/Ltac2/theories/Control.v @@ -0,0 +1,76 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* 'a := "ltac2" "throw". +(** Fatal exception throwing. This does not induce backtracking. *) + +(** Generic backtracking control *) + +Ltac2 @ external zero : exn -> 'a := "ltac2" "zero". +Ltac2 @ external plus : (unit -> 'a) -> (exn -> 'a) -> 'a := "ltac2" "plus". +Ltac2 @ external once : (unit -> 'a) -> 'a := "ltac2" "once". +Ltac2 @ external dispatch : (unit -> unit) list -> unit := "ltac2" "dispatch". +Ltac2 @ external extend : (unit -> unit) list -> (unit -> unit) -> (unit -> unit) list -> unit := "ltac2" "extend". +Ltac2 @ external enter : (unit -> unit) -> unit := "ltac2" "enter". +Ltac2 @ external case : (unit -> 'a) -> ('a * (exn -> 'a)) result := "ltac2" "case". + +(** Proof state manipulation *) + +Ltac2 @ external focus : int -> int -> (unit -> 'a) -> 'a := "ltac2" "focus". +Ltac2 @ external shelve : unit -> unit := "ltac2" "shelve". +Ltac2 @ external shelve_unifiable : unit -> unit := "ltac2" "shelve_unifiable". + +Ltac2 @ external new_goal : evar -> unit := "ltac2" "new_goal". +(** Adds the given evar to the list of goals as the last one. If it is + already defined in the current state, don't do anything. Panics if the + evar is not in the current state. *) + +Ltac2 @ external progress : (unit -> 'a) -> 'a := "ltac2" "progress". + +(** Goal inspection *) + +Ltac2 @ external goal : unit -> constr := "ltac2" "goal". +(** Panics if there is not exactly one goal under focus. Otherwise returns + the conclusion of this goal. *) + +Ltac2 @ external hyp : ident -> constr := "ltac2" "hyp". +(** Panics if there is more than one goal under focus. If there is no + goal under focus, looks for the section variable with the given name. + If there is one, looks for the hypothesis with the given name. *) + +Ltac2 @ external hyps : unit -> (ident * constr option * constr) list := "ltac2" "hyps". +(** Panics if there is more than one goal under focus. If there is no + goal under focus, returns the list of section variables. + If there is one, returns the list of hypotheses. In both cases, the + list is ordered with rightmost values being last introduced. *) + +(** Refinement *) + +Ltac2 @ external refine : (unit -> constr) -> unit := "ltac2" "refine". + +(** Evars *) + +Ltac2 @ external with_holes : (unit -> 'a) -> ('a -> 'b) -> 'b := "ltac2" "with_holes". +(** [with_holes x f] evaluates [x], then apply [f] to the result, and fails if + all evars generated by the call to [x] have not been solved when [f] + returns. *) + +(** Misc *) + +Ltac2 @ external time : string option -> (unit -> 'a) -> 'a := "ltac2" "time". +(** Displays the time taken by a tactic to evaluate. *) + +Ltac2 @ external abstract : ident option -> (unit -> unit) -> unit := "ltac2" "abstract". +(** Abstract a subgoal. *) + +Ltac2 @ external check_interrupt : unit -> unit := "ltac2" "check_interrupt". +(** For internal use. *) diff --git a/vendor/Ltac2/theories/Env.v b/vendor/Ltac2/theories/Env.v new file mode 100644 index 0000000000..c9b250f4ba --- /dev/null +++ b/vendor/Ltac2/theories/Env.v @@ -0,0 +1,27 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Std.reference option := "ltac2" "env_get". +(** Returns the global reference corresponding to the absolute name given as + argument if it exists. *) + +Ltac2 @ external expand : ident list -> Std.reference list := "ltac2" "env_expand". +(** Returns the list of all global references whose absolute name contains + the argument list as a prefix. *) + +Ltac2 @ external path : Std.reference -> ident list := "ltac2" "env_path". +(** Returns the absolute name of the given reference. Panics if the reference + does not exist. *) + +Ltac2 @ external instantiate : Std.reference -> constr := "ltac2" "env_instantiate". +(** Returns a fresh instance of the corresponding reference, in particular + generating fresh universe variables and constraints when this reference is + universe-polymorphic. *) diff --git a/vendor/Ltac2/theories/Fresh.v b/vendor/Ltac2/theories/Fresh.v new file mode 100644 index 0000000000..5e876bb077 --- /dev/null +++ b/vendor/Ltac2/theories/Fresh.v @@ -0,0 +1,26 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* t -> t := "ltac2" "fresh_free_union". + +Ltac2 @ external of_ids : ident list -> t := "ltac2" "fresh_free_of_ids". + +Ltac2 @ external of_constr : constr -> t := "ltac2" "fresh_free_of_constr". + +End Free. + +Ltac2 @ external fresh : Free.t -> ident -> ident := "ltac2" "fresh_fresh". +(** Generate a fresh identifier with the given base name which is not a + member of the provided set of free variables. *) diff --git a/vendor/Ltac2/theories/Ident.v b/vendor/Ltac2/theories/Ident.v new file mode 100644 index 0000000000..55456afbe2 --- /dev/null +++ b/vendor/Ltac2/theories/Ident.v @@ -0,0 +1,17 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* t -> bool := "ltac2" "ident_equal". + +Ltac2 @ external of_string : string -> t option := "ltac2" "ident_of_string". + +Ltac2 @ external to_string : t -> string := "ltac2" "ident_to_string". diff --git a/vendor/Ltac2/theories/Init.v b/vendor/Ltac2/theories/Init.v new file mode 100644 index 0000000000..16e7d7a6f9 --- /dev/null +++ b/vendor/Ltac2/theories/Init.v @@ -0,0 +1,69 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* int -> bool := "ltac2" "int_equal". +Ltac2 @ external compare : int -> int -> int := "ltac2" "int_compare". +Ltac2 @ external add : int -> int -> int := "ltac2" "int_add". +Ltac2 @ external sub : int -> int -> int := "ltac2" "int_sub". +Ltac2 @ external mul : int -> int -> int := "ltac2" "int_mul". +Ltac2 @ external neg : int -> int := "ltac2" "int_neg". diff --git a/vendor/Ltac2/theories/Ltac1.v b/vendor/Ltac2/theories/Ltac1.v new file mode 100644 index 0000000000..c4e0b606d0 --- /dev/null +++ b/vendor/Ltac2/theories/Ltac1.v @@ -0,0 +1,36 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* t := "ltac2" "ltac1_ref". +(** Returns the Ltac1 definition with the given absolute name. *) + +Ltac2 @ external run : t -> unit := "ltac2" "ltac1_run". +(** Runs an Ltac1 value, assuming it is a 'tactic', i.e. not returning + anything. *) + +Ltac2 @ external apply : t -> t list -> (t -> unit) -> unit := "ltac2" "ltac1_apply". +(** Applies an Ltac1 value to a list of arguments, and provides the result in + CPS style. It does **not** run the returned value. *) + +(** Conversion functions *) + +Ltac2 @ external of_constr : constr -> t := "ltac2" "ltac1_of_constr". +Ltac2 @ external to_constr : t -> constr option := "ltac2" "ltac1_to_constr". + +Ltac2 @ external of_list : t list -> t := "ltac2" "ltac1_of_list". +Ltac2 @ external to_list : t -> t list option := "ltac2" "ltac1_to_list". diff --git a/vendor/Ltac2/theories/Ltac2.v b/vendor/Ltac2/theories/Ltac2.v new file mode 100644 index 0000000000..ac90f63560 --- /dev/null +++ b/vendor/Ltac2/theories/Ltac2.v @@ -0,0 +1,24 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* unit := "ltac2" "print". + +Ltac2 @ external of_string : string -> message := "ltac2" "message_of_string". + +Ltac2 @ external of_int : int -> message := "ltac2" "message_of_int". + +Ltac2 @ external of_ident : ident -> message := "ltac2" "message_of_ident". + +Ltac2 @ external of_constr : constr -> message := "ltac2" "message_of_constr". +(** Panics if there is more than one goal under focus. *) + +Ltac2 @ external of_exn : exn -> message := "ltac2" "message_of_exn". +(** Panics if there is more than one goal under focus. *) + +Ltac2 @ external concat : message -> message -> message := "ltac2" "message_concat". diff --git a/vendor/Ltac2/theories/Notations.v b/vendor/Ltac2/theories/Notations.v new file mode 100644 index 0000000000..f4621656d6 --- /dev/null +++ b/vendor/Ltac2/theories/Notations.v @@ -0,0 +1,568 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* f e +| Val ans => + let (x, k) := ans in + Control.plus (fun _ => x) k +end. + +Ltac2 ifcatch t s f := +match Control.case t with +| Err e => f e +| Val ans => + let (x, k) := ans in + Control.plus (fun _ => s x) (fun e => s (k e)) +end. + +Ltac2 fail0 (_ : unit) := Control.enter (fun _ => Control.zero (Tactic_failure None)). + +Ltac2 Notation fail := fail0 (). + +Ltac2 try0 t := Control.enter (fun _ => orelse t (fun _ => ())). + +Ltac2 Notation try := try0. + +Ltac2 rec repeat0 (t : unit -> unit) := + Control.enter (fun () => + ifcatch (fun _ => Control.progress t) + (fun _ => Control.check_interrupt (); repeat0 t) (fun _ => ())). + +Ltac2 Notation repeat := repeat0. + +Ltac2 dispatch0 t (head, tail) := + match tail with + | None => Control.enter (fun _ => t (); Control.dispatch head) + | Some tacs => + let (def, rem) := tacs in + Control.enter (fun _ => t (); Control.extend head def rem) + end. + +Ltac2 Notation t(thunk(self)) ">" "[" l(dispatch) "]" : 4 := dispatch0 t l. + +Ltac2 do0 n t := + let rec aux n t := match Int.equal n 0 with + | true => () + | false => t (); aux (Int.sub n 1) t + end in + aux (n ()) t. + +Ltac2 Notation do := do0. + +Ltac2 Notation once := Control.once. + +Ltac2 progress0 tac := Control.enter (fun _ => Control.progress tac). + +Ltac2 Notation progress := progress0. + +Ltac2 rec first0 tacs := +match tacs with +| [] => Control.zero (Tactic_failure None) +| tac :: tacs => Control.enter (fun _ => orelse tac (fun _ => first0 tacs)) +end. + +Ltac2 Notation "first" "[" tacs(list0(thunk(tactic(6)), "|")) "]" := first0 tacs. + +Ltac2 complete tac := + let ans := tac () in + Control.enter (fun () => Control.zero (Tactic_failure None)); + ans. + +Ltac2 rec solve0 tacs := +match tacs with +| [] => Control.zero (Tactic_failure None) +| tac :: tacs => + Control.enter (fun _ => orelse (fun _ => complete tac) (fun _ => solve0 tacs)) +end. + +Ltac2 Notation "solve" "[" tacs(list0(thunk(tactic(6)), "|")) "]" := solve0 tacs. + +Ltac2 time0 tac := Control.time None tac. + +Ltac2 Notation time := time0. + +Ltac2 abstract0 tac := Control.abstract None tac. + +Ltac2 Notation abstract := abstract0. + +(** Base tactics *) + +(** Note that we redeclare notations that can be parsed as mere identifiers + as abbreviations, so that it allows to parse them as function arguments + without having to write them within parentheses. *) + +(** Enter and check evar resolution *) +Ltac2 enter_h ev f arg := +match ev with +| true => Control.enter (fun () => f ev (arg ())) +| false => + Control.enter (fun () => + Control.with_holes arg (fun x => f ev x)) +end. + +Ltac2 intros0 ev p := + Control.enter (fun () => Std.intros false p). + +Ltac2 Notation "intros" p(intropatterns) := intros0 false p. +Ltac2 Notation intros := intros. + +Ltac2 Notation "eintros" p(intropatterns) := intros0 true p. +Ltac2 Notation eintros := eintros. + +Ltac2 split0 ev bnd := + enter_h ev Std.split bnd. + +Ltac2 Notation "split" bnd(thunk(with_bindings)) := split0 false bnd. +Ltac2 Notation split := split. + +Ltac2 Notation "esplit" bnd(thunk(with_bindings)) := split0 true bnd. +Ltac2 Notation esplit := esplit. + +Ltac2 exists0 ev bnds := match bnds with +| [] => split0 ev (fun () => Std.NoBindings) +| _ => + let rec aux bnds := match bnds with + | [] => () + | bnd :: bnds => split0 ev bnd; aux bnds + end in + aux bnds +end. + +Ltac2 Notation "exists" bnd(list0(thunk(bindings), ",")) := exists0 false bnd. +(* Ltac2 Notation exists := exists. *) + +Ltac2 Notation "eexists" bnd(list0(thunk(bindings), ",")) := exists0 true bnd. +Ltac2 Notation eexists := eexists. + +Ltac2 left0 ev bnd := enter_h ev Std.left bnd. + +Ltac2 Notation "left" bnd(thunk(with_bindings)) := left0 false bnd. +Ltac2 Notation left := left. + +Ltac2 Notation "eleft" bnd(thunk(with_bindings)) := left0 true bnd. +Ltac2 Notation eleft := eleft. + +Ltac2 right0 ev bnd := enter_h ev Std.right bnd. + +Ltac2 Notation "right" bnd(thunk(with_bindings)) := right0 false bnd. +Ltac2 Notation right := right. + +Ltac2 Notation "eright" bnd(thunk(with_bindings)) := right0 true bnd. +Ltac2 Notation eright := eright. + +Ltac2 constructor0 ev n bnd := + enter_h ev (fun ev bnd => Std.constructor_n ev n bnd) bnd. + +Ltac2 Notation "constructor" := Control.enter (fun () => Std.constructor false). +Ltac2 Notation constructor := constructor. +Ltac2 Notation "constructor" n(tactic) bnd(thunk(with_bindings)) := constructor0 false n bnd. + +Ltac2 Notation "econstructor" := Control.enter (fun () => Std.constructor true). +Ltac2 Notation econstructor := econstructor. +Ltac2 Notation "econstructor" n(tactic) bnd(thunk(with_bindings)) := constructor0 true n bnd. + +Ltac2 specialize0 c pat := + enter_h false (fun _ c => Std.specialize c pat) c. + +Ltac2 Notation "specialize" c(thunk(seq(constr, with_bindings))) ipat(opt(seq("as", intropattern))) := + specialize0 c ipat. + +Ltac2 elim0 ev c bnd use := + let f ev (c, bnd, use) := Std.elim ev (c, bnd) use in + enter_h ev f (fun () => c (), bnd (), use ()). + +Ltac2 Notation "elim" c(thunk(constr)) bnd(thunk(with_bindings)) + use(thunk(opt(seq("using", constr, with_bindings)))) := + elim0 false c bnd use. + +Ltac2 Notation "eelim" c(thunk(constr)) bnd(thunk(with_bindings)) + use(thunk(opt(seq("using", constr, with_bindings)))) := + elim0 true c bnd use. + +Ltac2 apply0 adv ev cb cl := + Std.apply adv ev cb cl. + +Ltac2 Notation "eapply" + cb(list1(thunk(seq(constr, with_bindings)), ",")) + cl(opt(seq("in", ident, opt(seq("as", intropattern))))) := + apply0 true true cb cl. + +Ltac2 Notation "apply" + cb(list1(thunk(seq(constr, with_bindings)), ",")) + cl(opt(seq("in", ident, opt(seq("as", intropattern))))) := + apply0 true false cb cl. + +Ltac2 default_on_concl cl := +match cl with +| None => { Std.on_hyps := Some []; Std.on_concl := Std.AllOccurrences } +| Some cl => cl +end. + +Ltac2 pose0 ev p := + enter_h ev (fun ev (na, p) => Std.pose na p) p. + +Ltac2 Notation "pose" p(thunk(pose)) := + pose0 false p. + +Ltac2 Notation "epose" p(thunk(pose)) := + pose0 true p. + +Ltac2 Notation "set" p(thunk(pose)) cl(opt(clause)) := + Std.set false p (default_on_concl cl). + +Ltac2 Notation "eset" p(thunk(pose)) cl(opt(clause)) := + Std.set true p (default_on_concl cl). + +Ltac2 assert0 ev ast := + enter_h ev (fun _ ast => Std.assert ast) ast. + +Ltac2 Notation "assert" ast(thunk(assert)) := assert0 false ast. + +Ltac2 Notation "eassert" ast(thunk(assert)) := assert0 true ast. + +Ltac2 default_everywhere cl := +match cl with +| None => { Std.on_hyps := None; Std.on_concl := Std.AllOccurrences } +| Some cl => cl +end. + +Ltac2 Notation "remember" + c(thunk(open_constr)) + na(opt(seq("as", ident))) + pat(opt(seq("eqn", ":", intropattern))) + cl(opt(clause)) := + Std.remember false na c pat (default_everywhere cl). + +Ltac2 Notation "eremember" + c(thunk(open_constr)) + na(opt(seq("as", ident))) + pat(opt(seq("eqn", ":", intropattern))) + cl(opt(clause)) := + Std.remember true na c pat (default_everywhere cl). + +Ltac2 induction0 ev ic use := + let f ev use := Std.induction ev ic use in + enter_h ev f use. + +Ltac2 Notation "induction" + ic(list1(induction_clause, ",")) + use(thunk(opt(seq("using", constr, with_bindings)))) := + induction0 false ic use. + +Ltac2 Notation "einduction" + ic(list1(induction_clause, ",")) + use(thunk(opt(seq("using", constr, with_bindings)))) := + induction0 true ic use. + +Ltac2 generalize0 gen := + enter_h false (fun _ gen => Std.generalize gen) gen. + +Ltac2 Notation "generalize" + gen(thunk(list1(seq (open_constr, occurrences, opt(seq("as", ident))), ","))) := + generalize0 gen. + +Ltac2 destruct0 ev ic use := + let f ev use := Std.destruct ev ic use in + enter_h ev f use. + +Ltac2 Notation "destruct" + ic(list1(induction_clause, ",")) + use(thunk(opt(seq("using", constr, with_bindings)))) := + destruct0 false ic use. + +Ltac2 Notation "edestruct" + ic(list1(induction_clause, ",")) + use(thunk(opt(seq("using", constr, with_bindings)))) := + destruct0 true ic use. + +Ltac2 Notation "simple" "inversion" + arg(destruction_arg) + pat(opt(seq("as", intropattern))) + ids(opt(seq("in", list1(ident)))) := + Std.inversion Std.SimpleInversion arg pat ids. + +Ltac2 Notation "inversion" + arg(destruction_arg) + pat(opt(seq("as", intropattern))) + ids(opt(seq("in", list1(ident)))) := + Std.inversion Std.FullInversion arg pat ids. + +Ltac2 Notation "inversion_clear" + arg(destruction_arg) + pat(opt(seq("as", intropattern))) + ids(opt(seq("in", list1(ident)))) := + Std.inversion Std.FullInversionClear arg pat ids. + +Ltac2 Notation "red" cl(opt(clause)) := + Std.red (default_on_concl cl). +Ltac2 Notation red := red. + +Ltac2 Notation "hnf" cl(opt(clause)) := + Std.hnf (default_on_concl cl). +Ltac2 Notation hnf := hnf. + +Ltac2 Notation "simpl" s(strategy) pl(opt(seq(pattern, occurrences))) cl(opt(clause)) := + Std.simpl s pl (default_on_concl cl). +Ltac2 Notation simpl := simpl. + +Ltac2 Notation "cbv" s(strategy) cl(opt(clause)) := + Std.cbv s (default_on_concl cl). +Ltac2 Notation cbv := cbv. + +Ltac2 Notation "cbn" s(strategy) cl(opt(clause)) := + Std.cbn s (default_on_concl cl). +Ltac2 Notation cbn := cbn. + +Ltac2 Notation "lazy" s(strategy) cl(opt(clause)) := + Std.lazy s (default_on_concl cl). +Ltac2 Notation lazy := lazy. + +Ltac2 Notation "unfold" pl(list1(seq(reference, occurrences), ",")) cl(opt(clause)) := + Std.unfold pl (default_on_concl cl). + +Ltac2 fold0 pl cl := + let cl := default_on_concl cl in + Control.enter (fun () => Control.with_holes pl (fun pl => Std.fold pl cl)). + +Ltac2 Notation "fold" pl(thunk(list1(open_constr))) cl(opt(clause)) := + fold0 pl cl. + +Ltac2 Notation "pattern" pl(list1(seq(constr, occurrences), ",")) cl(opt(clause)) := + Std.pattern pl (default_on_concl cl). + +Ltac2 Notation "vm_compute" pl(opt(seq(pattern, occurrences))) cl(opt(clause)) := + Std.vm pl (default_on_concl cl). +Ltac2 Notation vm_compute := vm_compute. + +Ltac2 Notation "native_compute" pl(opt(seq(pattern, occurrences))) cl(opt(clause)) := + Std.native pl (default_on_concl cl). +Ltac2 Notation native_compute := native_compute. + +Ltac2 change0 p cl := + let (pat, c) := p in + Std.change pat c (default_on_concl cl). + +Ltac2 Notation "change" c(conversion) cl(opt(clause)) := change0 c cl. + +Ltac2 rewrite0 ev rw cl tac := + let cl := default_on_concl cl in + Std.rewrite ev rw cl tac. + +Ltac2 Notation "rewrite" + rw(list1(rewriting, ",")) + cl(opt(clause)) + tac(opt(seq("by", thunk(tactic)))) := + rewrite0 false rw cl tac. + +Ltac2 Notation "erewrite" + rw(list1(rewriting, ",")) + cl(opt(clause)) + tac(opt(seq("by", thunk(tactic)))) := + rewrite0 true rw cl tac. + +(** coretactics *) + +Ltac2 exact0 ev c := + Control.enter (fun _ => + match ev with + | true => + let c := c () in + Control.refine (fun _ => c) + | false => + Control.with_holes c (fun c => Control.refine (fun _ => c)) + end + ). + +Ltac2 Notation "exact" c(thunk(open_constr)) := exact0 false c. +Ltac2 Notation "eexact" c(thunk(open_constr)) := exact0 true c. + +Ltac2 Notation "intro" id(opt(ident)) mv(opt(move_location)) := Std.intro id mv. +Ltac2 Notation intro := intro. + +Ltac2 Notation "move" id(ident) mv(move_location) := Std.move id mv. + +Ltac2 Notation reflexivity := Std.reflexivity (). + +Ltac2 symmetry0 cl := + Std.symmetry (default_on_concl cl). + +Ltac2 Notation "symmetry" cl(opt(clause)) := symmetry0 cl. +Ltac2 Notation symmetry := symmetry. + +Ltac2 Notation "revert" ids(list1(ident)) := Std.revert ids. + +Ltac2 Notation assumption := Std.assumption (). + +Ltac2 Notation etransitivity := Std.etransitivity (). + +Ltac2 Notation admit := Std.admit (). + +Ltac2 clear0 ids := match ids with +| [] => Std.keep [] +| _ => Std.clear ids +end. + +Ltac2 Notation "clear" ids(list0(ident)) := clear0 ids. +Ltac2 Notation "clear" "-" ids(list1(ident)) := Std.keep ids. +Ltac2 Notation clear := clear. + +Ltac2 Notation refine := Control.refine. + +(** extratactics *) + +Ltac2 absurd0 c := Control.enter (fun _ => Std.absurd (c ())). + +Ltac2 Notation "absurd" c(thunk(open_constr)) := absurd0 c. + +Ltac2 subst0 ids := match ids with +| [] => Std.subst_all () +| _ => Std.subst ids +end. + +Ltac2 Notation "subst" ids(list0(ident)) := subst0 ids. +Ltac2 Notation subst := subst. + +Ltac2 Notation "discriminate" arg(opt(destruction_arg)) := + Std.discriminate false arg. +Ltac2 Notation discriminate := discriminate. + +Ltac2 Notation "ediscriminate" arg(opt(destruction_arg)) := + Std.discriminate true arg. +Ltac2 Notation ediscriminate := ediscriminate. + +Ltac2 Notation "injection" arg(opt(destruction_arg)) ipat(opt(seq("as", intropatterns))):= + Std.injection false ipat arg. + +Ltac2 Notation "einjection" arg(opt(destruction_arg)) ipat(opt(seq("as", intropatterns))):= + Std.injection true ipat arg. + +(** Auto *) + +Ltac2 default_db dbs := match dbs with +| None => Some [] +| Some dbs => + match dbs with + | None => None + | Some l => Some l + end +end. + +Ltac2 default_list use := match use with +| None => [] +| Some use => use +end. + +Ltac2 trivial0 use dbs := + let dbs := default_db dbs in + let use := default_list use in + Std.trivial Std.Off use dbs. + +Ltac2 Notation "trivial" + use(opt(seq("using", list1(thunk(constr), ",")))) + dbs(opt(seq("with", hintdb))) := trivial0 use dbs. + +Ltac2 Notation trivial := trivial. + +Ltac2 auto0 n use dbs := + let dbs := default_db dbs in + let use := default_list use in + Std.auto Std.Off n use dbs. + +Ltac2 Notation "auto" n(opt(tactic(0))) + use(opt(seq("using", list1(thunk(constr), ",")))) + dbs(opt(seq("with", hintdb))) := auto0 n use dbs. + +Ltac2 Notation auto := auto. + +Ltac2 new_eauto0 n use dbs := + let dbs := default_db dbs in + let use := default_list use in + Std.new_auto Std.Off n use dbs. + +Ltac2 Notation "new" "auto" n(opt(tactic(0))) + use(opt(seq("using", list1(thunk(constr), ",")))) + dbs(opt(seq("with", hintdb))) := new_eauto0 n use dbs. + +Ltac2 eauto0 n p use dbs := + let dbs := default_db dbs in + let use := default_list use in + Std.eauto Std.Off n p use dbs. + +Ltac2 Notation "eauto" n(opt(tactic(0))) p(opt(tactic(0))) + use(opt(seq("using", list1(thunk(constr), ",")))) + dbs(opt(seq("with", hintdb))) := eauto0 n p use dbs. + +Ltac2 Notation eauto := eauto. + +Ltac2 Notation "typeclasses_eauto" n(opt(tactic(0))) + dbs(opt(seq("with", list1(ident)))) := Std.typeclasses_eauto None n dbs. + +Ltac2 Notation "typeclasses_eauto" "bfs" n(opt(tactic(0))) + dbs(opt(seq("with", list1(ident)))) := Std.typeclasses_eauto (Some Std.BFS) n dbs. + +Ltac2 Notation typeclasses_eauto := typeclasses_eauto. + +(** Congruence *) + +Ltac2 f_equal0 () := ltac1:(f_equal). +Ltac2 Notation f_equal := f_equal0 (). + +(** Firstorder *) + +Ltac2 firstorder0 tac refs ids := + let refs := default_list refs in + let ids := default_list ids in + Std.firstorder tac refs ids. + +Ltac2 Notation "firstorder" + tac(opt(thunk(tactic))) + refs(opt(seq("using", list1(reference, ",")))) + ids(opt(seq("with", list1(ident)))) := firstorder0 tac refs ids. + +(** now *) + +Ltac2 now0 t := t (); ltac1:(easy). +Ltac2 Notation "now" t(thunk(self)) := now0 t. diff --git a/vendor/Ltac2/theories/Pattern.v b/vendor/Ltac2/theories/Pattern.v new file mode 100644 index 0000000000..8d1fb0cd8a --- /dev/null +++ b/vendor/Ltac2/theories/Pattern.v @@ -0,0 +1,145 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* context := + "ltac2" "pattern_empty_context". +(** A trivial context only made of the hole. *) + +Ltac2 @ external matches : t -> constr -> (ident * constr) list := + "ltac2" "pattern_matches". +(** If the term matches the pattern, returns the bound variables. If it doesn't, + fail with [Match_failure]. Panics if not focussed. *) + +Ltac2 @ external matches_subterm : t -> constr -> context * ((ident * constr) list) := + "ltac2" "pattern_matches_subterm". +(** Returns a stream of results corresponding to all of the subterms of the term + that matches the pattern as in [matches]. The stream is encoded as a + backtracking value whose last exception is [Match_failure]. The additional + value compared to [matches] is the context of the match, to be filled with + the instantiate function. *) + +Ltac2 @ external matches_vect : t -> constr -> constr array := + "ltac2" "pattern_matches_vect". +(** Internal version of [matches] that does not return the identifiers. *) + +Ltac2 @ external matches_subterm_vect : t -> constr -> context * constr array := + "ltac2" "pattern_matches_subterm_vect". +(** Internal version of [matches_subterms] that does not return the identifiers. *) + +Ltac2 @ external matches_goal : bool -> (match_kind * t) list -> (match_kind * t) -> + ident array * context array * constr array * context := + "ltac2" "pattern_matches_goal". +(** Given a list of patterns [hpats] for hypotheses and one pattern [cpat] for the + conclusion, [matches_goal rev hpats cpat] produces (a stream of) tuples of: + - An array of idents, whose size is the length of [hpats], corresponding to the + name of matched hypotheses. + - An array of contexts, whose size is the length of [hpats], corresponding to + the contexts matched for every hypothesis pattern. In case the match kind of + a hypothesis was [MatchPattern], the corresponding context is ensured to be empty. + - An array of terms, whose size is the total number of pattern variables without + duplicates. Terms are ordered by identifier order, e.g. ?a comes before ?b. + - A context corresponding to the conclusion, which is ensured to be empty if + the kind of [cpat] was [MatchPattern]. + This produces a backtracking stream of results containing all the possible + result combinations. The order of considered hypotheses is reversed if [rev] + is true. +*) + +Ltac2 @ external instantiate : context -> constr -> constr := + "ltac2" "pattern_instantiate". +(** Fill the hole of a context with the given term. *) + +(** Implementation of Ltac matching over terms and goals *) + +Ltac2 lazy_match0 t pats := + let rec interp m := match m with + | [] => Control.zero Match_failure + | p :: m => + let next _ := interp m in + let (knd, pat, f) := p in + let p := match knd with + | MatchPattern => + (fun _ => + let context := empty_context () in + let bind := matches_vect pat t in + fun _ => f context bind) + | MatchContext => + (fun _ => + let (context, bind) := matches_subterm_vect pat t in + fun _ => f context bind) + end in + Control.plus p next + end in + Control.once (fun () => interp pats) (). + +Ltac2 multi_match0 t pats := + let rec interp m := match m with + | [] => Control.zero Match_failure + | p :: m => + let next _ := interp m in + let (knd, pat, f) := p in + let p := match knd with + | MatchPattern => + (fun _ => + let context := empty_context () in + let bind := matches_vect pat t in + f context bind) + | MatchContext => + (fun _ => + let (context, bind) := matches_subterm_vect pat t in + f context bind) + end in + Control.plus p next + end in + interp pats. + +Ltac2 one_match0 t m := Control.once (fun _ => multi_match0 t m). + +Ltac2 lazy_goal_match0 rev pats := + let rec interp m := match m with + | [] => Control.zero Match_failure + | p :: m => + let next _ := interp m in + let (pat, f) := p in + let (phyps, pconcl) := pat in + let cur _ := + let (hids, hctx, subst, cctx) := matches_goal rev phyps pconcl in + fun _ => f hids hctx subst cctx + in + Control.plus cur next + end in + Control.once (fun () => interp pats) (). + +Ltac2 multi_goal_match0 rev pats := + let rec interp m := match m with + | [] => Control.zero Match_failure + | p :: m => + let next _ := interp m in + let (pat, f) := p in + let (phyps, pconcl) := pat in + let cur _ := + let (hids, hctx, subst, cctx) := matches_goal rev phyps pconcl in + f hids hctx subst cctx + in + Control.plus cur next + end in + interp pats. + +Ltac2 one_goal_match0 rev pats := Control.once (fun _ => multi_goal_match0 rev pats). diff --git a/vendor/Ltac2/theories/Std.v b/vendor/Ltac2/theories/Std.v new file mode 100644 index 0000000000..73b2ba02c4 --- /dev/null +++ b/vendor/Ltac2/theories/Std.v @@ -0,0 +1,263 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* constr), intro_pattern) +| IntroRewrite (bool) +] +with or_and_intro_pattern := [ +| IntroOrPattern (intro_pattern list list) +| IntroAndPattern (intro_pattern list) +]. + +Ltac2 Type destruction_arg := [ +| ElimOnConstr (unit -> constr_with_bindings) +| ElimOnIdent (ident) +| ElimOnAnonHyp (int) +]. + +Ltac2 Type induction_clause := { + indcl_arg : destruction_arg; + indcl_eqn : intro_pattern_naming option; + indcl_as : or_and_intro_pattern option; + indcl_in : clause option; +}. + +Ltac2 Type assertion := [ +| AssertType (intro_pattern option, constr, (unit -> unit) option) +| AssertValue (ident, constr) +]. + +Ltac2 Type repeat := [ +| Precisely (int) +| UpTo (int) +| RepeatStar +| RepeatPlus +]. + +Ltac2 Type orientation := [ LTR | RTL ]. + +Ltac2 Type rewriting := { + rew_orient : orientation option; + rew_repeat : repeat; + rew_equatn : (unit -> constr_with_bindings); +}. + +Ltac2 Type evar_flag := bool. +Ltac2 Type advanced_flag := bool. + +Ltac2 Type move_location := [ +| MoveAfter (ident) +| MoveBefore (ident) +| MoveFirst +| MoveLast +]. + +Ltac2 Type inversion_kind := [ +| SimpleInversion +| FullInversion +| FullInversionClear +]. + +(** Standard, built-in tactics. See Ltac1 for documentation. *) + +Ltac2 @ external intros : evar_flag -> intro_pattern list -> unit := "ltac2" "tac_intros". + +Ltac2 @ external apply : advanced_flag -> evar_flag -> + (unit -> constr_with_bindings) list -> (ident * (intro_pattern option)) option -> unit := "ltac2" "tac_apply". + +Ltac2 @ external elim : evar_flag -> constr_with_bindings -> constr_with_bindings option -> unit := "ltac2" "tac_elim". +Ltac2 @ external case : evar_flag -> constr_with_bindings -> unit := "ltac2" "tac_case". + +Ltac2 @ external generalize : (constr * occurrences * ident option) list -> unit := "ltac2" "tac_generalize". + +Ltac2 @ external assert : assertion -> unit := "ltac2" "tac_assert". +Ltac2 @ external enough : constr -> (unit -> unit) option option -> intro_pattern option -> unit := "ltac2" "tac_enough". + +Ltac2 @ external pose : ident option -> constr -> unit := "ltac2" "tac_pose". +Ltac2 @ external set : evar_flag -> (unit -> ident option * constr) -> clause -> unit := "ltac2" "tac_set". + +Ltac2 @ external remember : evar_flag -> ident option -> (unit -> constr) -> intro_pattern option -> clause -> unit := "ltac2" "tac_remember". + +Ltac2 @ external destruct : evar_flag -> induction_clause list -> + constr_with_bindings option -> unit := "ltac2" "tac_induction". + +Ltac2 @ external induction : evar_flag -> induction_clause list -> + constr_with_bindings option -> unit := "ltac2" "tac_induction". + +Ltac2 @ external red : clause -> unit := "ltac2" "tac_red". +Ltac2 @ external hnf : clause -> unit := "ltac2" "tac_hnf". +Ltac2 @ external simpl : red_flags -> (pattern * occurrences) option -> clause -> unit := "ltac2" "tac_simpl". +Ltac2 @ external cbv : red_flags -> clause -> unit := "ltac2" "tac_cbv". +Ltac2 @ external cbn : red_flags -> clause -> unit := "ltac2" "tac_cbn". +Ltac2 @ external lazy : red_flags -> clause -> unit := "ltac2" "tac_lazy". +Ltac2 @ external unfold : (reference * occurrences) list -> clause -> unit := "ltac2" "tac_unfold". +Ltac2 @ external fold : constr list -> clause -> unit := "ltac2" "tac_fold". +Ltac2 @ external pattern : (constr * occurrences) list -> clause -> unit := "ltac2" "tac_pattern". +Ltac2 @ external vm : (pattern * occurrences) option -> clause -> unit := "ltac2" "tac_vm". +Ltac2 @ external native : (pattern * occurrences) option -> clause -> unit := "ltac2" "tac_native". + +Ltac2 @ external eval_red : constr -> constr := "ltac2" "eval_red". +Ltac2 @ external eval_hnf : constr -> constr := "ltac2" "eval_hnf". +Ltac2 @ external eval_red : constr -> constr := "ltac2" "eval_red". +Ltac2 @ external eval_simpl : red_flags -> (pattern * occurrences) option -> constr -> constr := "ltac2" "eval_simpl". +Ltac2 @ external eval_cbv : red_flags -> constr -> constr := "ltac2" "eval_cbv". +Ltac2 @ external eval_cbn : red_flags -> constr -> constr := "ltac2" "eval_cbn". +Ltac2 @ external eval_lazy : red_flags -> constr -> constr := "ltac2" "eval_lazy". +Ltac2 @ external eval_unfold : (reference * occurrences) list -> constr -> constr := "ltac2" "eval_unfold". +Ltac2 @ external eval_fold : constr list -> constr -> constr := "ltac2" "eval_fold". +Ltac2 @ external eval_pattern : (constr * occurrences) list -> constr -> constr := "ltac2" "eval_pattern". +Ltac2 @ external eval_vm : (pattern * occurrences) option -> constr -> constr := "ltac2" "eval_vm". +Ltac2 @ external eval_native : (pattern * occurrences) option -> constr -> constr := "ltac2" "eval_native". + +Ltac2 @ external change : pattern option -> (constr array -> constr) -> clause -> unit := "ltac2" "tac_change". + +Ltac2 @ external rewrite : evar_flag -> rewriting list -> clause -> (unit -> unit) option -> unit := "ltac2" "tac_rewrite". + +Ltac2 @ external reflexivity : unit -> unit := "ltac2" "tac_reflexivity". + +Ltac2 @ external assumption : unit -> unit := "ltac2" "tac_assumption". + +Ltac2 @ external transitivity : constr -> unit := "ltac2" "tac_transitivity". + +Ltac2 @ external etransitivity : unit -> unit := "ltac2" "tac_etransitivity". + +Ltac2 @ external cut : constr -> unit := "ltac2" "tac_cut". + +Ltac2 @ external left : evar_flag -> bindings -> unit := "ltac2" "tac_left". +Ltac2 @ external right : evar_flag -> bindings -> unit := "ltac2" "tac_right". + +Ltac2 @ external constructor : evar_flag -> unit := "ltac2" "tac_constructor". +Ltac2 @ external split : evar_flag -> bindings -> unit := "ltac2" "tac_split". + +Ltac2 @ external constructor_n : evar_flag -> int -> bindings -> unit := "ltac2" "tac_constructorn". + +Ltac2 @ external intros_until : hypothesis -> unit := "ltac2" "tac_introsuntil". + +Ltac2 @ external symmetry : clause -> unit := "ltac2" "tac_symmetry". + +Ltac2 @ external rename : (ident * ident) list -> unit := "ltac2" "tac_rename". + +Ltac2 @ external revert : ident list -> unit := "ltac2" "tac_revert". + +Ltac2 @ external admit : unit -> unit := "ltac2" "tac_admit". + +Ltac2 @ external fix_ : ident option -> int -> unit := "ltac2" "tac_fix". +Ltac2 @ external cofix_ : ident option -> unit := "ltac2" "tac_cofix". + +Ltac2 @ external clear : ident list -> unit := "ltac2" "tac_clear". +Ltac2 @ external keep : ident list -> unit := "ltac2" "tac_keep". + +Ltac2 @ external clearbody : ident list -> unit := "ltac2" "tac_clearbody". + +Ltac2 @ external exact_no_check : constr -> unit := "ltac2" "tac_exactnocheck". +Ltac2 @ external vm_cast_no_check : constr -> unit := "ltac2" "tac_vmcastnocheck". +Ltac2 @ external native_cast_no_check : constr -> unit := "ltac2" "tac_nativecastnocheck". + +Ltac2 @ external inversion : inversion_kind -> destruction_arg -> intro_pattern option -> ident list option -> unit := "ltac2" "tac_inversion". + +(** coretactics *) + +Ltac2 @ external move : ident -> move_location -> unit := "ltac2" "tac_move". + +Ltac2 @ external intro : ident option -> move_location option -> unit := "ltac2" "tac_intro". + +Ltac2 @ external specialize : constr_with_bindings -> intro_pattern option -> unit := "ltac2" "tac_specialize". + +(** extratactics *) + +Ltac2 @ external discriminate : evar_flag -> destruction_arg option -> unit := "ltac2" "tac_discriminate". +Ltac2 @ external injection : evar_flag -> intro_pattern list option -> destruction_arg option -> unit := "ltac2" "tac_injection". + +Ltac2 @ external absurd : constr -> unit := "ltac2" "tac_absurd". +Ltac2 @ external contradiction : constr_with_bindings option -> unit := "ltac2" "tac_contradiction". + +Ltac2 @ external autorewrite : bool -> (unit -> unit) option -> ident list -> clause -> unit := "ltac2" "tac_autorewrite". + +Ltac2 @ external subst : ident list -> unit := "ltac2" "tac_subst". +Ltac2 @ external subst_all : unit -> unit := "ltac2" "tac_substall". + +(** auto *) + +Ltac2 Type debug := [ Off | Info | Debug ]. + +Ltac2 Type strategy := [ BFS | DFS ]. + +Ltac2 @ external trivial : debug -> (unit -> constr) list -> ident list option -> unit := "ltac2" "tac_trivial". + +Ltac2 @ external auto : debug -> int option -> (unit -> constr) list -> ident list option -> unit := "ltac2" "tac_auto". + +Ltac2 @ external new_auto : debug -> int option -> (unit -> constr) list -> ident list option -> unit := "ltac2" "tac_newauto". + +Ltac2 @ external eauto : debug -> int option -> int option -> (unit -> constr) list -> ident list option -> unit := "ltac2" "tac_eauto". + +Ltac2 @ external typeclasses_eauto : strategy option -> int option -> ident list option -> unit := "ltac2" "tac_typeclasses_eauto". + +(** firstorder *) + +Ltac2 @ external firstorder : (unit -> unit) option -> reference list -> ident list -> unit := "ltac2" "tac_firstorder". diff --git a/vendor/Ltac2/theories/String.v b/vendor/Ltac2/theories/String.v new file mode 100644 index 0000000000..99e1dab76b --- /dev/null +++ b/vendor/Ltac2/theories/String.v @@ -0,0 +1,14 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* char -> string := "ltac2" "string_make". +Ltac2 @external length : string -> int := "ltac2" "string_length". +Ltac2 @external get : string -> int -> char := "ltac2" "string_get". +Ltac2 @external set : string -> int -> char -> unit := "ltac2" "string_set". diff --git a/vendor/Ltac2/theories/dune b/vendor/Ltac2/theories/dune new file mode 100644 index 0000000000..1fe3ba28fe --- /dev/null +++ b/vendor/Ltac2/theories/dune @@ -0,0 +1,6 @@ +(coqlib + (name Ltac2) ; This determines the -R flag + (public_name ltac2.Ltac2) + (synopsis "Ltac 2 Plugin") + (libraries ltac2.plugin)) + -- cgit v1.2.3