diff options
| author | ppedrot | 2013-06-21 21:04:00 +0000 |
|---|---|---|
| committer | ppedrot | 2013-06-21 21:04:00 +0000 |
| commit | bd7da353ea503423206e329af7a56174cb39f435 (patch) | |
| tree | 275cce39ed6fb899660155a43ab0987c4f83025b /lib | |
| parent | 9024a91b59b9ecfb94e68b3748f2a9a66adcf515 (diff) | |
Splitted up Genarg in four different levels:
1. Genarg itself which only defines the abstract datatypes needed.
2. Genintern, first file of interp/, defining the intern and subst
functions.
3. Geninterp, first file of tactics/, defining the interp function.
4. Genprint, first file of printing/, dealing with the printers.
The Genarg file has no dependency and is in lib/, so that we can put
generic arguments everywhere, and in particular in ASTs.
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@16601 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'lib')
| -rw-r--r-- | lib/genarg.ml | 202 | ||||
| -rw-r--r-- | lib/genarg.mli | 250 | ||||
| -rw-r--r-- | lib/lib.mllib | 1 |
3 files changed, 453 insertions, 0 deletions
diff --git a/lib/genarg.ml b/lib/genarg.ml new file mode 100644 index 0000000000..dbde4652c7 --- /dev/null +++ b/lib/genarg.ml @@ -0,0 +1,202 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Pp +open Util + +type argument_type = + (* Basic types *) + | IntOrVarArgType + | IntroPatternArgType + | IdentArgType of bool + | VarArgType + | RefArgType + (* Specific types *) + | GenArgType + | SortArgType + | ConstrArgType + | ConstrMayEvalArgType + | QuantHypArgType + | OpenConstrArgType of bool + | ConstrWithBindingsArgType + | BindingsArgType + | RedExprArgType + | List0ArgType of argument_type + | List1ArgType of argument_type + | OptArgType of argument_type + | PairArgType of argument_type * argument_type + | ExtraArgType of string + +let rec argument_type_eq arg1 arg2 = match arg1, arg2 with +| IntOrVarArgType, IntOrVarArgType -> true +| IntroPatternArgType, IntroPatternArgType -> true +| IdentArgType b1, IdentArgType b2 -> (b1 : bool) == b2 +| VarArgType, VarArgType -> true +| RefArgType, RefArgType -> true +| GenArgType, GenArgType -> true +| SortArgType, SortArgType -> true +| ConstrArgType, ConstrArgType -> true +| ConstrMayEvalArgType, ConstrMayEvalArgType -> true +| QuantHypArgType, QuantHypArgType -> true +| OpenConstrArgType b1, OpenConstrArgType b2 -> (b1 : bool) == b2 +| ConstrWithBindingsArgType, ConstrWithBindingsArgType -> true +| BindingsArgType, BindingsArgType -> true +| RedExprArgType, RedExprArgType -> true +| List0ArgType arg1, List0ArgType arg2 -> argument_type_eq arg1 arg2 +| List1ArgType arg1, List1ArgType arg2 -> argument_type_eq arg1 arg2 +| OptArgType arg1, OptArgType arg2 -> argument_type_eq arg1 arg2 +| PairArgType (arg1l, arg1r), PairArgType (arg2l, arg2r) -> + argument_type_eq arg1l arg2l && argument_type_eq arg1r arg2r +| ExtraArgType s1, ExtraArgType s2 -> CString.equal s1 s2 +| _ -> false + +type ('raw, 'glob, 'top) genarg_type = argument_type + +type 'a uniform_genarg_type = ('a, 'a, 'a) genarg_type +(** Alias for concision *) + +(* Dynamics but tagged by a type expression *) + +type rlevel +type glevel +type tlevel + +type 'a generic_argument = argument_type * Obj.t +type raw_generic_argument = rlevel generic_argument +type glob_generic_argument = glevel generic_argument +type typed_generic_argument = tlevel generic_argument + +let rawwit t = t +let glbwit t = t +let topwit t = t + +let wit_list0 t = List0ArgType t + +let wit_list1 t = List1ArgType t + +let wit_opt t = OptArgType t + +let wit_pair t1 t2 = PairArgType (t1,t2) + +let in_gen t o = (t,Obj.repr o) +let out_gen t (t',o) = if argument_type_eq t t' then Obj.magic o else failwith "out_gen" +let genarg_tag (s,_) = s + +let fold_list0 f = function + | (List0ArgType t, l) -> + List.fold_right (fun x -> f (in_gen t x)) (Obj.magic l) + | _ -> failwith "Genarg: not a list0" + +let fold_list1 f = function + | (List1ArgType t, l) -> + List.fold_right (fun x -> f (in_gen t x)) (Obj.magic l) + | _ -> failwith "Genarg: not a list1" + +let fold_opt f a = function + | (OptArgType t, l) -> + (match Obj.magic l with + | None -> a + | Some x -> f (in_gen t x)) + | _ -> failwith "Genarg: not a opt" + +let fold_pair f = function + | (PairArgType (t1,t2), l) -> + let (x1,x2) = Obj.magic l in + f (in_gen t1 x1) (in_gen t2 x2) + | _ -> failwith "Genarg: not a pair" + +let app_list0 f = function + | (List0ArgType t as u, l) -> + let o = Obj.magic l in + (u, Obj.repr (List.map (fun x -> out_gen t (f (in_gen t x))) o)) + | _ -> failwith "Genarg: not a list0" + +let app_list1 f = function + | (List1ArgType t as u, l) -> + let o = Obj.magic l in + (u, Obj.repr (List.map (fun x -> out_gen t (f (in_gen t x))) o)) + | _ -> failwith "Genarg: not a list1" + +let app_opt f = function + | (OptArgType t as u, l) -> + let o = Obj.magic l in + (u, Obj.repr (Option.map (fun x -> out_gen t (f (in_gen t x))) o)) + | _ -> failwith "Genarg: not an opt" + +let app_pair f1 f2 = function + | (PairArgType (t1,t2) as u, l) -> + let (o1,o2) = Obj.magic l in + let o1 = out_gen t1 (f1 (in_gen t1 o1)) in + let o2 = out_gen t2 (f2 (in_gen t2 o2)) in + (u, Obj.repr (o1,o2)) + | _ -> failwith "Genarg: not a pair" + +let has_type (t, v) u = argument_type_eq t u + +let unquote x = x + +type an_arg_of_this_type = Obj.t + +let in_generic t x = (t, Obj.repr x) + +type ('a,'b) abstract_argument_type = argument_type +type 'a raw_abstract_argument_type = ('a,rlevel) abstract_argument_type +type 'a glob_abstract_argument_type = ('a,glevel) abstract_argument_type +type 'a typed_abstract_argument_type = ('a,tlevel) abstract_argument_type + +(** Creating args *) + +let (arg0_map : Obj.t option String.Map.t ref) = ref String.Map.empty + +let create_arg opt name = + if String.Map.mem name !arg0_map then + Errors.anomaly (str "generic argument already declared: " ++ str name) + else + let () = arg0_map := String.Map.add name (Obj.magic opt) !arg0_map in + ExtraArgType name + +let make0 = create_arg + +let default_empty_value t = + let rec aux = function + | List0ArgType _ -> Some (Obj.repr []) + | OptArgType _ -> Some (Obj.repr None) + | PairArgType(t1, t2) -> + (match aux t1, aux t2 with + | Some v1, Some v2 -> Some (Obj.repr (v1, v2)) + | _ -> None) + | ExtraArgType s -> + String.Map.find s !arg0_map + | _ -> None in + match aux t with + | Some v -> Some (Obj.obj v) + | None -> None + +(** Hackish part *) + +let arg0_names = ref (String.Map.empty : string String.Map.t) +(** We use this table to associate a name to a given witness, to use it with + the extension mechanism. This is REALLY ad-hoc, but I do not know how to + do so nicely either. *) + +let register_name0 t name = match t with +| ExtraArgType s -> + let () = assert (not (String.Map.mem s !arg0_names)) in + arg0_names := String.Map.add s name !arg0_names +| _ -> failwith "register_name0" + +let get_name0 name = + String.Map.find name !arg0_names + +module Unsafe = +struct + +let inj tpe x = (tpe, x) +let prj (_, x) = x + +end diff --git a/lib/genarg.mli b/lib/genarg.mli new file mode 100644 index 0000000000..14133fff03 --- /dev/null +++ b/lib/genarg.mli @@ -0,0 +1,250 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +(** The route of a generic argument, from parsing to evaluation. +In the following diagram, "object" can be tactic_expr, constr, tactic_arg, etc. + +{% \begin{%}verbatim{% }%} + parsing in_raw out_raw + char stream ---> raw_object ---> raw_object generic_argument -------+ + encapsulation decaps| + | + V + raw_object + | + globalization | + V + glob_object + | + encaps | + in_glob | + V + glob_object generic_argument + | + out in out_glob | + object <--- object generic_argument <--- object <--- glob_object <---+ + | decaps encaps interp decaps + | + V +effective use +{% \end{%}verbatim{% }%} + +To distinguish between the uninterpreted (raw), globalized and +interpreted worlds, we annotate the type [generic_argument] by a +phantom argument which is either [constr_expr], [glob_constr] or +[constr]. + +Transformation for each type : +{% \begin{%}verbatim{% }%} +tag raw open type cooked closed type + +BoolArgType bool bool +IntArgType int int +IntOrVarArgType int or_var int +StringArgType string (parsed w/ "") string +PreIdentArgType string (parsed w/o "") (vernac only) +IdentArgType true identifier identifier +IdentArgType false identifier (pattern_ident) identifier +IntroPatternArgType intro_pattern_expr intro_pattern_expr +VarArgType identifier located identifier +RefArgType reference global_reference +QuantHypArgType quantified_hypothesis quantified_hypothesis +ConstrArgType constr_expr constr +ConstrMayEvalArgType constr_expr may_eval constr +OpenConstrArgType open_constr_expr open_constr +ConstrWithBindingsArgType constr_expr with_bindings constr with_bindings +BindingsArgType constr_expr bindings constr bindings +List0ArgType of argument_type +List1ArgType of argument_type +OptArgType of argument_type +ExtraArgType of string '_a '_b +{% \end{%}verbatim{% }%} +*) + +(** {5 Generic types} *) + +type ('raw, 'glob, 'top) genarg_type +(** Generic types. ['raw] is the OCaml lowest level, ['glob] is the globalized + one, and ['top] the internalized one. *) + +type 'a uniform_genarg_type = ('a, 'a, 'a) genarg_type +(** Alias for concision when the three types agree. *) + +val make0 : 'raw option -> string -> ('raw, 'glob, 'top) genarg_type +(** Create a new generic type of argument: force to associate + unique ML types at each of the three levels. *) + +val create_arg : 'raw option -> string -> ('raw, 'glob, 'top) genarg_type +(** Alias for [make0]. *) + +(** {5 Specialized types} *) + +(** All of [rlevel], [glevel] and [tlevel] must be non convertible + to ensure the injectivity of the type inference from type + ['co generic_argument] to [('a,'co) abstract_argument_type]; + this guarantees that, for 'co fixed, the type of + out_gen is monomorphic over 'a, hence type-safe +*) + +type rlevel +type glevel +type tlevel + +type ('a, 'co) abstract_argument_type +(** Type at level ['co] represented by an OCaml value of type ['a]. *) + +type 'a raw_abstract_argument_type = ('a, rlevel) abstract_argument_type +(** Specialized type at raw level. *) + +type 'a glob_abstract_argument_type = ('a, glevel) abstract_argument_type +(** Specialized type at globalized level. *) + +type 'a typed_abstract_argument_type = ('a, tlevel) abstract_argument_type +(** Specialized type at internalized level. *) + +(** {6 Projections} *) + +val rawwit : ('a, 'b, 'c) genarg_type -> ('a, rlevel) abstract_argument_type +(** Projection on the raw type constructor. *) + +val glbwit : ('a, 'b, 'c) genarg_type -> ('b, glevel) abstract_argument_type +(** Projection on the globalized type constructor. *) + +val topwit : ('a, 'b, 'c) genarg_type -> ('c, tlevel) abstract_argument_type +(** Projection on the internalized type constructor. *) + +(** {5 Generic arguments} *) + +type 'a generic_argument +(** A inhabitant of ['level generic_argument] is a inhabitant of some type at + level ['level], together with the representation of this type. *) + +type raw_generic_argument = rlevel generic_argument +type glob_generic_argument = glevel generic_argument +type typed_generic_argument = tlevel generic_argument + +(** {6 Constructors} *) + +val in_gen : ('a, 'co) abstract_argument_type -> 'a -> 'co generic_argument +(** [in_gen t x] embeds an argument of type [t] into a generic argument. *) + +val out_gen : ('a, 'co) abstract_argument_type -> 'co generic_argument -> 'a +(** [out_gen t x] recovers an argument of type [t] from a generic argument. It + fails if [x] has not the right dynamic type. *) + +val has_type : 'co generic_argument -> ('a, 'co) abstract_argument_type -> bool +(** [has_type v t] tells whether [v] has type [t]. If true, it ensures that + [out_gen t v] will not raise a dynamic type exception. *) + +(** {6 Manipulation of generic arguments} + +Those functions fail if they are applied to an argument which has not the right +dynamic type. *) + +val fold_list0 : + ('a generic_argument -> 'c -> 'c) -> 'a generic_argument -> 'c -> 'c + +val fold_list1 : + ('a generic_argument -> 'c -> 'c) -> 'a generic_argument -> 'c -> 'c + +val fold_opt : + ('a generic_argument -> 'c) -> 'c -> 'a generic_argument -> 'c + +val fold_pair : + ('a generic_argument -> 'a generic_argument -> 'c) -> + 'a generic_argument -> 'c + +(** [app_list0] fails if applied to an argument not of tag [List0 t] + for some [t]; it's the responsability of the caller to ensure it *) + +val app_list0 : ('a generic_argument -> 'b generic_argument) -> +'a generic_argument -> 'b generic_argument + +val app_list1 : ('a generic_argument -> 'b generic_argument) -> +'a generic_argument -> 'b generic_argument + +val app_opt : ('a generic_argument -> 'b generic_argument) -> +'a generic_argument -> 'b generic_argument + +val app_pair : + ('a generic_argument -> 'b generic_argument) -> + ('a generic_argument -> 'b generic_argument) + -> 'a generic_argument -> 'b generic_argument + +(** {6 Type reification} *) + +type argument_type = + (** Basic types *) + | IntOrVarArgType + | IntroPatternArgType + | IdentArgType of bool + | VarArgType + | RefArgType + (** Specific types *) + | GenArgType + | SortArgType + | ConstrArgType + | ConstrMayEvalArgType + | QuantHypArgType + | OpenConstrArgType of bool + | ConstrWithBindingsArgType + | BindingsArgType + | RedExprArgType + | List0ArgType of argument_type + | List1ArgType of argument_type + | OptArgType of argument_type + | PairArgType of argument_type * argument_type + | ExtraArgType of string + +val argument_type_eq : argument_type -> argument_type -> bool + +val genarg_tag : 'a generic_argument -> argument_type + +val unquote : ('a, 'co) abstract_argument_type -> argument_type + +(** {5 Basic generic type constructors} *) + +(** {6 Parameterized types} *) + +val wit_list0 : ('a, 'b, 'c) genarg_type -> ('a list, 'b list, 'c list) genarg_type +val wit_list1 : ('a, 'b, 'c) genarg_type -> ('a list, 'b list, 'c list) genarg_type +val wit_opt : ('a, 'b, 'c) genarg_type -> ('a option, 'b option, 'c option) genarg_type +val wit_pair : ('a1, 'b1, 'c1) genarg_type -> ('a2, 'b2, 'c2) genarg_type -> + ('a1 * 'a2, 'b1 * 'b2, 'c1 * 'c2) genarg_type + +(** {5 Magic used by the parser} *) + +val default_empty_value : ('raw, 'glb, 'top) genarg_type -> 'raw option + +val register_name0 : ('a, 'b, 'c) genarg_type -> string -> unit +(** Used by the extension to give a name to types. The string should be the + absolute path of the argument witness, e.g. + [register_name0 wit_toto "MyArg.wit_toto"]. *) + +val get_name0 : string -> string +(** Return the absolute path of a given witness. *) + +(** {5 Unsafe loophole} *) + +module Unsafe : +sig + +(** Unsafe magic functions. Not for kids. This is provided here as a loophole to + escape this module. Do NOT use outside of the dedicated areas. NOT. EVER. *) + +val inj : argument_type -> Obj.t -> 'lev generic_argument +(** Injects an object as generic argument. !!!BEWARE!!! only do this as + [inj tpe x] where: + + 1. [tpe] is the reification of a [('a, 'b, 'c) genarg_type]; + 2. [x] has type ['a], ['b] or ['c] according to the return level ['lev]. *) + +val prj : 'lev generic_argument -> Obj.t +(** Recover the contents of a generic argument. *) + +end diff --git a/lib/lib.mllib b/lib/lib.mllib index be27eed7d8..f28b49a4fe 100644 --- a/lib/lib.mllib +++ b/lib/lib.mllib @@ -16,3 +16,4 @@ Rtree Heap Dnet Unionfind +Genarg |
