diff options
author | Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr> | 2014-08-29 17:05:13 +0200 |
---|---|---|
committer | Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr> | 2014-08-29 18:43:06 +0200 |
commit | dac4d8952c5fc234f5b6245e39a73c2ca07555ee (patch) | |
tree | 6a5da7a1a6b3d02183c3e786e84b991bff72a171 /lib | |
parent | e3f2781feb49325615affaaef2853c56874e6c22 (diff) |
Type-safe version of genarg list / pair / opt functions.
Diffstat (limited to 'lib')
-rw-r--r-- | lib/genarg.ml | 81 | ||||
-rw-r--r-- | lib/genarg.mli | 41 |
2 files changed, 40 insertions, 82 deletions
diff --git a/lib/genarg.ml b/lib/genarg.ml index b85fef087..ce0041036 100644 --- a/lib/genarg.ml +++ b/lib/genarg.ml @@ -96,60 +96,6 @@ 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_list f = function - | (ListArgType t, l) -> - List.fold_right (fun x -> f (in_gen t x)) (Obj.magic l) - | _ -> failwith "Genarg: not a list" - -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_list f = function - | (ListArgType 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_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" - -module Monadic (M:Monad.S) = struct - - let app_list f = function - | (ListArgType t as u, l) -> - let o = Obj.magic l in - let open M in - let apply x = - f (in_gen t x) >>= fun y -> - return (out_gen t y) - in - M.List.map apply o >>= fun r -> - return (u, Obj.repr r) - | _ -> failwith "Genarg: not a list0" - -end - let has_type (t, v) u = argument_type_eq t u let unquote x = x @@ -170,6 +116,33 @@ type ('r, 'l) unpacker = let unpack pack (t, obj) = pack.unpacker t (Obj.obj obj) +(** Type transformers *) + +type ('r, 'l) list_unpacker = + { list_unpacker : 'a 'b 'c. ('a, 'b, 'c) genarg_type -> + ('a list, 'b list, 'c list, 'l) cast -> 'r } + +let list_unpack pack (t, obj) = match t with +| ListArgType t -> pack.list_unpacker t (Obj.obj obj) +| _ -> failwith "out_gen" + +type ('r, 'l) opt_unpacker = + { opt_unpacker : 'a 'b 'c. ('a, 'b, 'c) genarg_type -> + ('a option, 'b option, 'c option, 'l) cast -> 'r } + +let opt_unpack pack (t, obj) = match t with +| OptArgType t -> pack.opt_unpacker t (Obj.obj obj) +| _ -> failwith "out_gen" + +type ('r, 'l) pair_unpacker = + { pair_unpacker : 'a1 'a2 'b1 'b2 'c1 'c2. + ('a1, 'b1, 'c1) genarg_type -> ('a2, 'b2, 'c2) genarg_type -> + (('a1 * 'a2), ('b1 * 'b2), ('c1 * 'c2), 'l) cast -> 'r } + +let pair_unpack pack (t, obj) = match t with +| PairArgType (t1, t2) -> pack.pair_unpacker t1 t2 (Obj.obj obj) +| _ -> failwith "out_gen" + (** Creating args *) let (arg0_map : Obj.t option String.Map.t ref) = ref String.Map.empty diff --git a/lib/genarg.mli b/lib/genarg.mli index e96f21c10..90e701239 100644 --- a/lib/genarg.mli +++ b/lib/genarg.mli @@ -160,39 +160,24 @@ val unpack : ('r, 'l) unpacker -> 'l generic_argument -> 'r Those functions fail if they are applied to an argument which has not the right dynamic type. *) -val fold_list : - ('a generic_argument -> 'c -> 'c) -> 'a generic_argument -> 'c -> 'c +type ('r, 'l) list_unpacker = + { list_unpacker : 'a 'b 'c. ('a, 'b, 'c) genarg_type -> + ('a list, 'b list, 'c list, 'l) cast -> 'r } -val fold_opt : - ('a generic_argument -> 'c) -> 'c -> 'a generic_argument -> 'c +val list_unpack : ('r, 'l) list_unpacker -> 'l generic_argument -> 'r -val fold_pair : - ('a generic_argument -> 'a generic_argument -> 'c) -> - 'a generic_argument -> 'c +type ('r, 'l) opt_unpacker = + { opt_unpacker : 'a 'b 'c. ('a, 'b, 'c) genarg_type -> + ('a option, 'b option, 'c option, 'l) cast -> 'r } -(** [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 opt_unpack : ('r, 'l) opt_unpacker -> 'l generic_argument -> 'r -val app_list : ('a generic_argument -> 'b generic_argument) -> -'a generic_argument -> 'b generic_argument +type ('r, 'l) pair_unpacker = + { pair_unpacker : 'a1 'a2 'b1 'b2 'c1 'c2. + ('a1, 'b1, 'c1) genarg_type -> ('a2, 'b2, 'c2) genarg_type -> + (('a1 * 'a2), ('b1 * 'b2), ('c1 * 'c2), 'l) cast -> 'r } -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 - -module Monadic (M:Monad.S) : sig - - (** [Monadic.app_list f x] maps the monadic computation [f] on - elements of [x], provided [x] has the tag [List0 t] for some [t]. It - fails otherwise. *) - val app_list : ('a generic_argument -> 'b generic_argument M.t) -> - 'a generic_argument -> 'b generic_argument M.t - -end +val pair_unpack : ('r, 'l) pair_unpacker -> 'l generic_argument -> 'r (** {6 Type reification} *) |