Quot/Nominal/Test.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 23 Feb 2010 13:32:35 +0100
changeset 1223 160343d86a6f
parent 1194 3d54fcc5f41a
child 1228 c179ad9d2446
permissions -rw-r--r--
"raw"-ified the term-constructors and types given in the specification

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 *}

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:  term list 
          ^^^^^^^^^
          the equations of the binding functions
          (Trueprop equations)
*}

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))

val constr_parser =
  P.binding -- Scan.repeat anno_typ

(* datatype parser *)
val dt_parser =
  ((P.type_args -- P.binding -- P.opt_infix) >> P.triple1) -- 
    (P.$$$ "=" |-- P.enum1 "|" ((constr_parser -- bind_parser -- P.opt_mixfix) >> P.triple_swap))

(* function equation parser *)
val fun_parser = 
  Scan.optional (P.$$$ "binder" |-- P.fixes -- SpecParse.where_alt_specs) ([],[])

(* main parser *)
val main_parser =
  P.and_list1 dt_parser -- fun_parser

end
*}

(* adds "_raw" to the end of constants and types *)
ML {*
fun add_raw s = s ^ "_raw"
fun raw_bind bn = Binding.suffix_name "_raw" bn

fun raw_str [] s = s
  | raw_str (s'::ss) s = 
      if (Long_Name.base_name s) = s' 
      then add_raw s
      else raw_str ss s

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_aterm trm_strs (Const (a, T)) = Const (raw_str trm_strs a, T)
  | raw_aterm trm_strs (Free (a, T)) = Free (raw_str trm_strs a, T)
  | raw_aterm trm_strs trm = trm
 
fun raw_term trm_strs ty_strs trm =
  trm |> Term.map_aterms (raw_aterm trm_strs) |> map_types (raw_typ ty_strs) 

fun raw_dts ty_strs dts =
let
  fun raw_dts_aux1 (bind, tys, mx) =
    (raw_bind bind, map (raw_typ ty_strs) tys, mx)

  fun raw_dts_aux2 (ty_args, bind, mx, constrs) =
    (ty_args, raw_bind bind, mx, map raw_dts_aux1 constrs)
in
  map raw_dts_aux2 dts
end

fun get_constr_strs dts =
  flat (map (fn (_, _, _, constrs) => map (fn (bn, _, _) => Binding.name_of bn) constrs) dts)

fun get_bn_fun_strs bn_funs =
  map (fn (bn_fun, _, _) => Binding.name_of bn_fun) bn_funs
*}

ML {*
fun nominal_dt_decl dt_names dts bn_funs bn_eqs lthy =
let
  val conf = Datatype.default_config
  
  val dt_names' = map add_raw dt_names
  val dts' = raw_dts dt_names dts
  val constr_strs = get_constr_strs dts
  val bn_fun_strs = get_bn_fun_strs bn_funs
  
  val bn_funs' = map (fn (bn, opt_ty, mx) => 
    (raw_bind bn, Option.map (raw_typ dt_names) opt_ty, mx)) bn_funs
  val bn_eqs' = map (fn trm => 
    (Attrib.empty_binding, raw_term (constr_strs @ bn_fun_strs) dt_names trm)) bn_eqs
in
  lthy
  |> (Local_Theory.theory_result (Datatype.add_datatype conf dt_names' dts'))
  ||>> Primrec.add_primrec bn_funs' bn_eqs'
end 
*}

local_setup {*
let
  val T0 = Type ("Test.foo", [])
  val T1 = T0 --> @{typ "nat"}
in
nominal_dt_decl
  ["foo"]
  [([], @{binding "foo"}, NoSyn, 
    [(@{binding "myzero"}, [], NoSyn),(@{binding "mysuc"}, [Type ("Test.foo", [])], NoSyn)])]
  [(@{binding "Bar"}, SOME T1, NoSyn)]
  [HOLogic.mk_Trueprop (HOLogic.mk_eq 
        (Free ("Bar", T1) $ Const ("Test.foo_raw.myzero", T0), @{term "0::nat"})),
   HOLogic.mk_Trueprop (HOLogic.mk_eq 
        (Free ("Bar", T1) $ (Const ("Test.foo_raw.mysuc", T0 --> T0) $ Free ("n", T0)), @{term "0::nat"}))]
  #> snd
end
*}

typ foo_raw
thm foo_raw.induct
term myzero_raw
term mysuc_raw
thm Bar_raw.simps

(* 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


(* probably can be done with a clever morphism trick *)