theory Test
imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp"
begin
atom_decl name
(* tests *)
ML {*
Datatype.datatype_cmd;
Datatype.add_datatype Datatype.default_config;
Primrec.add_primrec_cmd;
Primrec.add_primrec;
Primrec.add_primrec_simple;
*}
section {* test for setting up a primrec on the ML-level *}
datatype simple = Foo
local_setup {*
Primrec.add_primrec [(@{binding "Bar"}, SOME @{typ "simple \<Rightarrow> nat"}, NoSyn)]
[(Attrib.empty_binding, HOLogic.mk_Trueprop @{term "Bar Foo = Suc 0"})] #> snd
*}
thm Bar.simps
lemma "Bar Foo = XXX"
apply(simp)
oops
section {* test for setting up a datatype on the ML-level *}
setup {*
Datatype.add_datatype Datatype.default_config
["foo"]
[([], @{binding "foo"}, NoSyn,
[(@{binding "zero"}, [], NoSyn),(@{binding "suc"}, [Type ("Test.foo", [])], NoSyn)])
] #> snd
*}
ML {*
PolyML.makestring (#descr (Datatype.the_info @{theory} "Test.foo"))
*}
thm foo.recs
thm foo.induct
(* adds "_raw" to the end of constants and types *)
ML {*
fun raw_str [] s = s
| raw_str (s'::ss) s = (if s = s' then s ^ "_str" else raw_str ss s)
val raw_strs = map raw_str
fun raw_typ ty_strs (Type (a, Ts)) = Type (raw_str ty_strs a, map (raw_typ ty_strs) Ts)
| raw_typ ty_strs T = T
fun raw_term trm_strs ty_strs (Const (a, T)) = Const (raw_str trm_strs a, raw_typ ty_strs T)
| raw_term trm_strs ty_strs (Abs (a, T, t)) = Abs (a, T, raw_term trm_strs ty_strs t)
| raw_term trm_strs ty_strs (t $ u) = (raw_term trm_strs ty_strs t) $ (raw_term trm_strs ty_strs u)
fun raw_binding bn = Binding.suffix_name "_raw" bn
*}
section{* Interface for nominal_datatype *}
text {*
Nominal-Datatype-part:
1st Arg: string list
^^^^^^^^^^^
strings of the types to be defined
2nd Arg: (string list * binding * mixfix * (binding * typ list * mixfix) list) list
^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^ ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^
type(s) to be defined constructors list
(ty args, name, syn) (name, typs, syn)
Binder-Function-part:
3rd Arg: (binding * typ option * mixfix) list
^^^^^^^^^^^^^^^^^^^^^^^^^^^^^
binding function(s)
to be defined
(name, type, syn)
4th Arg: (attrib * term) list
^^^^^^^^^^^^^^^^^^^
the equations of the binding functions
(Trueprop equations)
*}
ML {* Datatype.add_datatype *}
ML {*
fun raw_definitions dt_names dts bn_funs bn_eqs lthy =
let
val conf = Datatype.default_config
val bn_funs' = map (fn (bn, ty) => (raw_binding bn, ty, NoSyn)) bn_funs
val bn_eqs' = map (pair Attrib.empty_binding) bn_eqs
in
lthy
|> (Local_Theory.theory_result (Datatype.add_datatype conf dt_names dts))
||>> Primrec.add_primrec bn_funs' bn_eqs'
end
*}
ML {* @{binding "zero"} *}
local_setup {*
raw_definitions
["foo'"]
[([], @{binding "foo'"}, NoSyn,
[(@{binding "zero'"}, [], NoSyn),(@{binding "suc'"}, [Type ("Test.foo'", [])], NoSyn)])]
[(@{binding "Bar'"}, SOME @{typ "simple \<Rightarrow> nat"})]
[HOLogic.mk_Trueprop @{term "Bar'_raw Foo = Suc 0"}]
#> snd
*}
typ foo'
thm Bar'.simps
text {*****************************************************}
ML {*
(* nominal datatype parser *)
local
structure P = OuterParse
in
val _ = OuterKeyword.keyword "bind"
val anno_typ = Scan.option (P.name --| P.$$$ "::") -- P.typ
(* binding specification *)
(* should use and_list *)
val bind_parser =
P.enum "," ((P.$$$ "bind" |-- P.term) -- (P.$$$ "in" |-- P.name))
(* datatype parser *)
val dt_parser =
P.type_args -- P.binding -- P.opt_infix --
(P.$$$ "=" |-- P.enum1 "|" (P.binding -- (Scan.repeat anno_typ) -- bind_parser -- P.opt_mixfix))
(* function equation parser *)
val fun_parser =
Scan.optional (P.$$$ "binder" |-- P.fixes -- SpecParse.where_alt_specs) ([],[])
(* main parser *)
val parser =
P.and_list1 dt_parser -- fun_parser
end
*}
(* printing stuff *)
ML {*
fun print_str_option NONE = "NONE"
| print_str_option (SOME s) = (s:bstring)
fun print_anno_typ lthy (NONE, ty) = Syntax.string_of_typ lthy ty
| print_anno_typ lthy (SOME x, ty) = x ^ ":" ^ Syntax.string_of_typ lthy ty
fun print_bind (s1, s2) = "bind " ^ s1 ^ " in " ^ s2
fun print_constr lthy (((name, anno_tys), bds), syn) =
(Binding.name_of name ^ " of " ^ (commas (map (print_anno_typ lthy) anno_tys)) ^ " "
^ (commas (map print_bind bds)) ^ " "
^ (Pretty.str_of (Syntax.pretty_mixfix syn)))
fun print_single_dt lthy (((s2, s3), syn), constrs) =
["constructor declaration"]
@ ["type arguments: "] @ s2
@ ["datatype name: ", Binding.name_of s3]
@ ["typ mixfix: ", Pretty.string_of (Syntax.pretty_mixfix syn)]
@ ["constructors: "] @ (map (print_constr lthy) constrs)
|> separate "\n"
|> implode
fun print_single_fun_eq lthy (at, s) =
["function equations: ", (Syntax.string_of_term lthy s)]
|> separate "\n"
|> implode
fun print_single_fun_fix lthy (b, _, _) =
["functions: ", Binding.name_of b]
|> separate "\n"
|> implode
fun dt_name (((s2, s3), syn), constrs) = Binding.name_of s3
fun print_dts (dts, (funs, feqs)) lthy =
let
val _ = warning (implode (separate "\n" (map (print_single_dt lthy) dts)));
val _ = warning (implode (separate "\n" (map (print_single_fun_fix lthy) funs)));
val _ = warning (implode (separate "\n" (map (print_single_fun_eq lthy) feqs)));
in
()
end
*}
ML {*
fun transp_ty_arg (anno, ty) = ty
fun transp_constr (((constr_name, ty_args), bind), mx) =
(constr_name, map transp_ty_arg ty_args, mx)
fun transp_dt (((ty_args, ty_name), mx), constrs) =
(ty_args, ty_name, mx, map transp_constr constrs)
*}
ML {*
fun dt_defn dts thy =
let
val names = map (fn (_, s, _, _) => Binding.name_of s) dts
val thy' = Datatype.datatype_cmd names dts thy
in
thy'
end
*}
ML {*
fun fn_defn [] [] lthy = lthy
| fn_defn funs feqs lthy =
Function_Fun.add_fun Function_Common.default_config funs feqs false lthy
fun fn_defn_cmd [] [] lthy = lthy
| fn_defn_cmd funs feqs lthy =
Function_Fun.add_fun_cmd Function_Common.default_config funs feqs false lthy
*}
ML {*
fun parse_fun lthy (b, NONE, m) = (b, NONE, m)
| parse_fun lthy (b, SOME ty, m) = (b, SOME (Syntax.read_typ lthy ty), m)
fun parse_feq lthy (b, t) = (b, Syntax.read_prop lthy t);
fun parse_anno_ty lthy (anno, ty) = (anno, Syntax.read_typ lthy ty);
fun parse_constr lthy (((constr_name, ty_args), bind), mx) =
(((constr_name, map (parse_anno_ty lthy) ty_args), bind), mx);
fun parse_dt lthy (((ty_args, ty_name), mx), constrs) =
(((ty_args, ty_name), mx), map (parse_constr lthy) constrs);
*}
ML {*
fun nominal_datatype2_cmd (dt_strs, (fun_strs, feq_strs)) thy =
let
val thy' = dt_defn (map transp_dt dt_strs) thy
val lthy = Theory_Target.init NONE thy'
val lthy' = fn_defn_cmd fun_strs feq_strs lthy
val funs = map (parse_fun lthy') fun_strs
val feqs = map (parse_feq lthy') feq_strs
val dts = map (parse_dt lthy') dt_strs
val _ = print_dts (dts, (funs, feqs)) lthy'
in
Local_Theory.exit_global lthy'
end
*}
(* Command Keyword *)
ML {*
let
val kind = OuterKeyword.thy_decl
in
OuterSyntax.command "nominal_datatype" "test" kind
(parser >> (Toplevel.theory o nominal_datatype2_cmd))
end
*}
text {* example 1 *}
nominal_datatype lam =
VAR "name"
| APP "lam" "lam"
| LET bp::"bp" t::"lam" bind "bi bp" in t ("Let _ in _" [100,100] 100)
and bp =
BP "name" "lam" ("_ ::= _" [100,100] 100)
binder
bi::"bp \<Rightarrow> name set"
where
"bi (BP x t) = {x}"
print_theorems
text {* examples 2 *}
nominal_datatype trm =
Var "name"
| App "trm" "trm"
| Lam x::"name" t::"trm" bind x in t
| Let p::"pat" "trm" t::"trm" bind "f p" in t
and pat =
PN
| PS "name"
| PD "name \<times> name"
binder
f::"pat \<Rightarrow> name set"
where
"f PN = {}"
| "f (PS x) = {x}"
| "f (PD (x,y)) = {x,y}"
nominal_datatype trm2 =
Var2 "name"
| App2 "trm2" "trm2"
| Lam2 x::"name" t::"trm2" bind x in t
| Let2 p::"pat2" "trm2" t::"trm2" bind "f2 p" in t
and pat2 =
PN2
| PS2 "name"
| PD2 "pat2 \<times> pat2"
binder
f2::"pat2 \<Rightarrow> name set"
where
"f2 PN2 = {}"
| "f2 (PS2 x) = {x}"
| "f2 (PD2 (p1, p2)) = (f2 p1) \<union> (f2 p2)"
text {* example type schemes *}
datatype ty =
Var "name"
| Fun "ty" "ty"
nominal_datatype tyS =
All xs::"name list" ty::"ty" bind xs in ty
end